got Christian changes
authorChengsong
Wed, 13 Jul 2022 08:35:45 +0100
changeset 566 94604a5fd271
parent 565 0497408a3598 (current diff)
parent 563 c92a41d9c4da (diff)
child 567 28cb8089ec36
got Christian changes
--- a/thys/BitCoded.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys/BitCoded.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -29,7 +29,7 @@
 where
   "decode' ds ZERO = (Void, [])"
 | "decode' ds ONE = (Void, ds)"
-| "decode' ds (CHAR d) = (Char d, ds)"
+| "decode' ds (CH d) = (Char d, ds)"
 | "decode' [] (ALT r1 r2) = (Void, [])"
 | "decode' (Z # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r1 in (Left v, ds'))"
 | "decode' (S # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r2 in (Right v, ds'))"
@@ -111,7 +111,7 @@
 where
   "erase AZERO = ZERO"
 | "erase (AONE _) = ONE"
-| "erase (ACHAR _ c) = CHAR c"
+| "erase (ACHAR _ c) = CH c"
 | "erase (AALTs _ []) = ZERO"
 | "erase (AALTs _ [r]) = (erase r)"
 | "erase (AALTs bs (r#rs)) = ALT (erase r) (erase (AALTs bs rs))"
@@ -165,7 +165,7 @@
 fun intern :: "rexp \<Rightarrow> arexp" where
   "intern ZERO = AZERO"
 | "intern ONE = AONE []"
-| "intern (CHAR c) = ACHAR [] c"
+| "intern (CH c) = ACHAR [] c"
 | "intern (ALT r1 r2) = AALT [] (fuse [Z] (intern r1)) 
                                 (fuse [S]  (intern r2))"
 | "intern (SEQ r1 r2) = ASEQ [] (intern r1) (intern r2)"
--- a/thys3/Paper.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/Paper.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -32,6 +32,7 @@
   ALT ("_ + _" [77,77] 78) and
   SEQ ("_ \<cdot> _" [77,77] 78) and
   STAR ("_\<^sup>*" [79] 78) and
+  NTIMES ("_\<^sup>{\<^sup>_\<^sup>}" [79] 78) and
 
   val.Void ("Empty" 78) and
   val.Char ("Char _" [1000] 78) and
@@ -94,7 +95,13 @@
 Brzozowski's derivatives \cite{Brzozowski1964} is that they are neatly
 expressible in any functional language, and easily definable and
 reasoned about in theorem provers---the definitions just consist of
-inductive datatypes and simple recursive functions.  Derivatives of a
+inductive datatypes and simple recursive functions. Another neat
+feature of derivatives is that they can be easily extended to bounded
+regular expressions, such as @{term "NTIMES r n"}, where numbers or
+intervals specify how many times a regular expression should be used
+during matching.
+
+Derivatives of a
 regular expression, written @{term "der c r"}, give a simple solution
 to the problem of matching a string @{term s} with a regular
 expression @{term r}: if the derivative of @{term r} w.r.t.\ (in
@@ -231,7 +238,8 @@
   @{term "CH c"} $\mid$
   @{term "ALT r\<^sub>1 r\<^sub>2"} $\mid$
   @{term "SEQ r\<^sub>1 r\<^sub>2"} $\mid$
-  @{term "STAR r"} 
+  @{term "STAR r"} $\mid$
+  @{term "NTIMES r n"}
   \end{center}
 
   \noindent where @{const ZERO} stands for the regular expression that does
@@ -239,7 +247,13 @@
   only the empty string and @{term c} for matching a character literal.
   The constructors $+$ and $\cdot$ represent alternatives and sequences, respectively.
   We sometimes omit the $\cdot$ in a sequence regular expression for brevity. 
-  The
+  In our work here we also add to the usual ``basic'' regular expressions 
+  the bounded regular expression @{term "NTIMES r n"} where the @{term n}
+  specifies that @{term r} should match exactly @{term n}-times. For
+  brevity we omit the other bounded regular expressions
+  @{text "r"}$^{\{..n\}}$, @{text "r"}$^{\{n..\}}$ and @{text "r"}$^{\{n..m\}}$
+  which specify an interval for how many times @{text r} should match. Our
+  results extend straightforwardly also to them. The
   \emph{language} of a regular expression, written $L(r)$, is defined as usual
   and we omit giving the definition here (see for example \cite{AusafDyckhoffUrban2016}).
 
--- a/thys3/ROOT	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/ROOT	Wed Jul 13 08:35:45 2022 +0100
@@ -3,8 +3,8 @@
      "HOL-Library.Sublist"
      "RegLangs"
      "PosixSpec"
-     "Positions"
-     "PDerivs"
+     (*"Positions"*)
+     (*"PDerivs"*)
      "Lexer"
      "LexerSimp"
      "Blexer"
--- a/thys3/src/BasicIdentities.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/src/BasicIdentities.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -9,6 +9,7 @@
 | RSEQ rrexp rrexp
 | RALTS "rrexp list"
 | RSTAR rrexp
+| RNTIMES rrexp nat
 
 abbreviation
   "RALT r1 r2 \<equiv> RALTS [r1, r2]"
@@ -23,7 +24,7 @@
 | "rnullable (RALTS rs) = (\<exists>r \<in> set rs. rnullable r)"
 | "rnullable (RSEQ r1 r2) = (rnullable r1 \<and> rnullable r2)"
 | "rnullable (RSTAR r) = True"
-
+| "rnullable (RNTIMES r n) = (if n = 0 then True else rnullable r)"
 
 fun
  rder :: "char \<Rightarrow> rrexp \<Rightarrow> rrexp"
@@ -36,8 +37,8 @@
      (if rnullable r1
       then RALT   (RSEQ (rder c r1) r2) (rder c r2)
       else RSEQ   (rder c r1) r2)"
-| "rder c (RSTAR r) = RSEQ  (rder c r) (RSTAR r)"
-
+| "rder c (RSTAR r) = RSEQ  (rder c r) (RSTAR r)"   
+| "rder c (RNTIMES r n) = (if n = 0 then RZERO else RSEQ (rder c r) (RNTIMES r (n - 1)))"
 
 fun 
   rders :: "rrexp \<Rightarrow> string \<Rightarrow> rrexp"
@@ -191,6 +192,7 @@
 | "rsize (RALTS  rs) = Suc (sum_list (map rsize rs))"
 | "rsize (RSEQ  r1 r2) = Suc (rsize r1 + rsize r2)"
 | "rsize (RSTAR  r) = Suc (rsize r)"
+| "rsize (RNTIMES  r n) = Suc (rsize r) + n"
 
 abbreviation rsizes where
   "rsizes rs \<equiv> sum_list (map rsize rs)"
@@ -237,7 +239,9 @@
      apply(case_tac r2)
         apply simp_all
      apply(case_tac r2)
-  apply simp_all
+         apply simp_all
+  apply(case_tac r2)
+         apply simp_all
   done
 
 lemma ralts_cap_mono:
@@ -329,7 +333,9 @@
    apply(case_tac r2)
   apply simp_all
   apply(case_tac r2)
-       apply simp_all
+         apply simp_all
+apply(case_tac r2)
+         apply simp_all
   done
 
 lemma rders__onechar:
@@ -374,7 +380,7 @@
 | "good (RSEQ  _ RZERO) = False"
 | "good (RSEQ r1 r2) = (good r1 \<and> good r2)"
 | "good (RSTAR r) = True"
-
+| "good (RNTIMES r n) = True"
 
 lemma  k0a:
   shows "rflts [RALTS rs] =   rs"
@@ -424,9 +430,8 @@
      apply (metis flts1 good.simps(5) good.simps(6) k0a neq_Nil_conv)
     apply (metis flts1 good.simps(5) good.simps(6) k0a neq_Nil_conv)
    apply fastforce
-  apply(simp)
-  done  
-
+   apply(simp)
+  by simp
 
 
 lemma flts3:
@@ -458,7 +463,9 @@
   apply(case_tac r2)
         apply(simp_all)
   apply(case_tac r2)
-       apply(simp_all)
+         apply(simp_all)
+apply(case_tac r2)
+         apply(simp_all)
   done
 
 lemma rsize0:
@@ -509,7 +516,7 @@
   apply(induct rs )
    apply(auto)
     apply (metis list.set_intros(1) nn1qq nonalt.elims(3))
-  apply (metis nonalt.elims(2) nonnested.simps(3) nonnested.simps(4) nonnested.simps(5) nonnested.simps(6) nonnested.simps(7))
+  apply (metis nonalt.elims(2) nonnested.simps(3) nonnested.simps(4) nonnested.simps(5) nonnested.simps(6) nonnested.simps(7) nonnested.simps(8))
   using bbbbs1 apply fastforce
   by (metis bbbbs1 list.set_intros(2) nn1qq)
 
@@ -552,8 +559,8 @@
   apply(case_tac "\<exists>bs. rsimp r1 = RONE")
     apply(auto)[1]
   using idiot apply fastforce
-  using idiot2 nonnested.simps(11) apply presburger
-  by (metis (mono_tags, lifting) Diff_empty image_iff list.set_map nn1bb nn1c rdistinct_set_equality1)
+  apply (simp add: idiot2)
+  by (metis (mono_tags, lifting) image_iff list.set_map nn1bb nn1c rdistinct_set_equality)
 
 lemma nonalt_flts_rd:
   shows "\<lbrakk>xa \<in> set (rdistinct (rflts (map rsimp rs)) {})\<rbrakk>
@@ -603,6 +610,7 @@
    apply (simp add: elem_smaller_than_set)
   by (metis Diff_empty flts3 rdistinct_set_equality1)
 
+thm Diff_empty flts3 rdistinct_set_equality1
   
 lemma good1:
   shows "good (rsimp a) \<or> rsimp a = RZERO"
@@ -632,8 +640,9 @@
   apply blast
   apply fastforce
   using less_add_Suc2 apply blast  
-  using less_iff_Suc_add by blast
-
+  using less_iff_Suc_add apply blast
+  using good.simps(45) rsimp.simps(7) by presburger
+  
 
 
 fun
@@ -645,28 +654,51 @@
 | "RL (RSEQ r1 r2) = (RL r1) ;; (RL r2)"
 | "RL (RALTS rs) = (\<Union> (set (map RL rs)))"
 | "RL (RSTAR r) = (RL r)\<star>"
+| "RL (RNTIMES r n) = (RL r) ^^ n"
 
+lemma pow_rempty_iff:
+  shows "[] \<in> (RL r) ^^ n \<longleftrightarrow> (if n = 0 then True else [] \<in> (RL r))"
+  by (induct n) (auto simp add: Sequ_def)
 
 lemma RL_rnullable:
   shows "rnullable r = ([] \<in> RL r)"
   apply(induct r)
-  apply(auto simp add: Sequ_def)
+        apply(auto simp add: Sequ_def pow_rempty_iff)
   done
 
+lemma concI_if_Nil1: "[] \<in> A \<Longrightarrow> xs : B \<Longrightarrow> xs \<in> A ;; B"
+by (metis append_Nil concI)
+
+
+lemma empty_pow_add:
+  fixes A::"string set"
+  assumes "[] \<in> A" "s \<in> A ^^ n"
+  shows "s \<in> A ^^ (n + m)"
+  using assms
+  apply(induct m arbitrary: n)
+   apply(auto simp add: Sequ_def)
+  done
+
+(*
+lemma der_pow:
+  shows "Der c (A ^^ n) = (if n = 0 then {} else (Der c A) ;; (A ^^ (n - 1)))"
+  apply(induct n arbitrary: A)
+   apply(auto)
+  by (smt (verit, best) Suc_pred concE concI concI_if_Nil2 conc_pow_comm lang_pow.simps(2))
+*)
+
 lemma RL_rder:
   shows "RL (rder c r) = Der c (RL r)"
   apply(induct r)
-  apply(auto simp add: Sequ_def Der_def)
+  apply(auto simp add: Sequ_def Der_def)[5]
         apply (metis append_Cons)
   using RL_rnullable apply blast
   apply (metis append_eq_Cons_conv)
   apply (metis append_Cons)
-  apply (metis RL_rnullable append_eq_Cons_conv)
-  apply (metis Star.step append_Cons)
-  using Star_decomp by auto
-
-
-
+    apply (metis RL_rnullable append_eq_Cons_conv)
+  apply simp
+  apply(simp)
+  done
 
 lemma RL_rsimp_RSEQ:
   shows "RL (rsimp_SEQ r1 r2) = (RL r1 ;; RL r2)"
@@ -843,12 +875,14 @@
   apply(case_tac "rsimp aa")
   apply simp+
   apply (metis no_alt_short_list_after_simp no_further_dB_after_simp)
-  by simp
+   apply(simp)
+  apply(simp)
+  done
 
 lemma identity_wwo0:
   shows "rsimp (rsimp_ALTs (RZERO # rs)) = rsimp (rsimp_ALTs rs)"
-  by (metis idem_after_simp1 list.exhaust list.simps(8) list.simps(9) rflts.simps(2) rsimp.simps(2) rsimp.simps(3) rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3))
-
+  apply (metis idem_after_simp1 list.exhaust list.simps(8) list.simps(9) rflts.simps(2) rsimp.simps(2) rsimp.simps(3) rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3))
+  done
 
 lemma distinct_removes_last:
   shows "\<lbrakk>a \<in> set as\<rbrakk>
@@ -994,12 +1028,15 @@
   
   apply(subgoal_tac "\<forall>r \<in> set( rflts (map rsimp rsa)). good r")
    apply(case_tac "rdistinct (rflts (map rsimp rsa)) {}")
-    apply simp
-   apply(case_tac "listb")
-    apply simp+
-  apply (metis Cons_eq_appendI good1_flatten rflts.simps(3) rsimp.simps(2) rsimp_ALTs.simps(3))
-  by (metis (mono_tags, lifting) flts3 good1 image_iff list.set_map)
-
+     apply simp
+  apply auto[1]
+  apply simp
+  apply(simp)
+   apply(case_tac "lista")
+  apply simp_all
+ 
+   apply (metis append_Cons append_Nil good1_flatten rflts.simps(2) rsimp.simps(2) rsimp_ALTs.elims)
+   by (metis (no_types, opaque_lifting) append_Cons append_Nil good1_flatten rflts.simps(2) rsimp.simps(2) rsimp_ALTs.elims)
 
 lemma last_elem_out:
   shows "\<lbrakk>x \<notin> set xs; x \<notin> rset \<rbrakk> \<Longrightarrow> rdistinct (xs @ [x]) rset = rdistinct xs rset @ [x]"
@@ -1127,8 +1164,9 @@
   apply fastforce
   apply fastforce
   apply fastforce
-  by fastforce
-
+  apply fastforce
+  by simp
+  
 
 lemma distinct_removes_duplicate_flts:
   shows " a \<in> set rsa
@@ -1143,29 +1181,19 @@
       apply(subgoal_tac " rdistinct (rflts (map rsimp rsa @ [rsimp a])) {} =  
                           rdistinct (rflts (map rsimp rsa @ [RONE])) {}")
       apply (simp only:)
-       apply(subst flts_keeps1)
-  apply (metis distinct_removes_last(1) rflts_def_idiot2 rrexp.simps(20) rrexp.simps(6))
+        apply(subst flts_keeps1)
+  apply (metis distinct_removes_last(1) flts_append in_set_conv_decomp rflts.simps(4))
       apply presburger
         apply(subgoal_tac " rdistinct (rflts (map rsimp rsa @ [rsimp a]))    {} =  
                             rdistinct ((rflts (map rsimp rsa)) @ [RCHAR x]) {}")
       apply (simp only:)
-      prefer 2
-      apply (metis flts_keeps_others rrexp.distinct(21) rrexp.distinct(3))
-  apply (metis distinct_removes_last(1) rflts_def_idiot2 rrexp.distinct(21) rrexp.distinct(3))
-
-    apply (metis distinct_removes_last(1) flts_keeps_others rflts_def_idiot2 rrexp.distinct(25) rrexp.distinct(5))
-   prefer 2
-   apply (metis distinct_removes_last(1) flts_keeps_others flts_removes0 rflts_def_idiot2 rrexp.distinct(29))
-  apply(subgoal_tac "rflts (map rsimp rsa @ [rsimp a]) = rflts (map rsimp rsa) @ x")
-  prefer 2
-  apply (simp add: rflts_spills_last)
-  apply(subgoal_tac "\<forall> r \<in> set x. r \<in> set (rflts (map rsimp rsa))")
-    prefer 2
-  apply (metis (mono_tags, lifting) image_iff image_set spilled_alts_contained)
-  apply (metis rflts_spills_last)
-  by (metis distinct_removes_list spilled_alts_contained)
-
-
+       prefer 2
+       apply (metis flts_append rflts.simps(1) rflts.simps(5))
+  apply (metis distinct_removes_last(1) rflts_def_idiot2 rrexp.distinct(25) rrexp.distinct(3))
+  apply (metis distinct_removes_last(1) flts_append rflts.simps(1) rflts.simps(6) rflts_def_idiot2 rrexp.distinct(31) rrexp.distinct(5))
+  apply (metis distinct_removes_list rflts_spills_last spilled_alts_contained)
+  apply (metis distinct_removes_last(1) flts_append good.simps(1) good.simps(44) rflts.simps(1) rflts.simps(7) rflts_def_idiot2 rrexp.distinct(37))
+  by (metis distinct_removes_last(1) flts_append rflts.simps(1) rflts.simps(8) rflts_def_idiot2 rrexp.distinct(11) rrexp.distinct(39))
 
 (*some basic facts about rsimp*)
 
--- a/thys3/src/Blexer.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/src/Blexer.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -1,6 +1,6 @@
 
 theory Blexer
-  imports "Lexer" "PDerivs"
+  imports "Lexer"
 begin
 
 section \<open>Bit-Encodings\<close>
@@ -17,13 +17,22 @@
 | "code (Stars []) = [S]"
 | "code (Stars (v # vs)) =  (Z # code v) @ code (Stars vs)"
 
+fun sz where
+  "sz ZERO = 0"
+| "sz ONE = 0"
+| "sz (CH _) = 0"
+| "sz (SEQ r1 r2) = 1 + sz r1 + sz r2"
+| "sz (ALT r1 r2) = 1 + sz r1 + sz r2"
+| "sz (STAR r) = 1 + sz r"
+| "sz (NTIMES r n) = 1 + n + sz r"
+
 
 fun 
   Stars_add :: "val \<Rightarrow> val \<Rightarrow> val"
 where
   "Stars_add v (Stars vs) = Stars (v # vs)"
 
-function
+function (sequential)
   decode' :: "bit list \<Rightarrow> rexp \<Rightarrow> (val * bit list)"
 where
   "decode' bs ZERO = (undefined, bs)"
@@ -39,6 +48,12 @@
 | "decode' (Z # bs) (STAR r) = (let (v, bs') = decode' bs r in
                                     let (vs, bs'') = decode' bs' (STAR r) 
                                     in (Stars_add v vs, bs''))"
+| "decode' [] (NTIMES r n) = (Void, [])"
+| "decode' (S # bs) (NTIMES r n) = (Stars [], bs)"
+(*| "decode' (Z # bs) (NTIMES r 0) = (undefined, bs)"*)
+| "decode' (Z # bs) (NTIMES r n) = (let (v, bs') = decode' bs r in
+                                    let (vs, bs'') = decode' bs' (NTIMES r (n - 1)) 
+                                    in (Stars_add v vs, bs''))"
 by pat_completeness auto
 
 lemma decode'_smaller:
@@ -48,12 +63,15 @@
 apply(induct bs r)
 apply(auto simp add: decode'.psimps split: prod.split)
 using dual_order.trans apply blast
-by (meson dual_order.trans le_SucI)
+apply (meson dual_order.trans le_SucI)
+  apply (meson le_SucI le_trans)
+  done
 
 termination "decode'"  
-apply(relation "inv_image (measure(%cs. size cs) <*lex*> measure(%s. size s)) (%(ds,r). (r,ds))") 
+apply(relation "inv_image (measure(%cs. sz cs) <*lex*> measure(%s. size s)) (%(ds,r). (r,ds))") 
 apply(auto dest!: decode'_smaller)
-by (metis less_Suc_eq_le snd_conv)
+   apply (metis less_Suc_eq_le snd_conv)
+  by (metis less_Suc_eq_le snd_conv)
 
 definition
   decode :: "bit list \<Rightarrow> rexp \<Rightarrow> val option"
@@ -69,13 +87,23 @@
   apply(auto)
   done
 
+lemma decode'_code_NTIMES:
+  assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> (\<forall>x. decode' (code v @ x) r = (v, x))" 
+  shows "decode' (code (Stars vs) @ ds) (NTIMES r n) = (Stars vs, ds)"
+  using assms
+  apply(induct vs arbitrary: n r ds)
+   apply(auto)
+  done
+
+
 lemma decode'_code:
   assumes "\<Turnstile> v : r"
   shows "decode' ((code v) @ ds) r = (v, ds)"
 using assms
   apply(induct v r arbitrary: ds) 
   apply(auto)
-  using decode'_code_Stars by blast
+  using decode'_code_Stars apply blast
+   by (metis Un_iff decode'_code_NTIMES set_append)  
 
 lemma decode_code:
   assumes "\<Turnstile> v : r"
@@ -93,6 +121,7 @@
 | ASEQ "bit list" arexp arexp
 | AALTs "bit list" "arexp list"
 | ASTAR "bit list" arexp
+| ANTIMES "bit list" arexp nat
 
 abbreviation
   "AALT bs r1 r2 \<equiv> AALTs bs [r1, r2]"
@@ -104,6 +133,7 @@
 | "asize (AALTs cs rs) = Suc (sum_list (map asize rs))"
 | "asize (ASEQ cs r1 r2) = Suc (asize r1 + asize r2)"
 | "asize (ASTAR cs r) = Suc (asize r)"
+| "asize (ANTIMES cs r n) = Suc (asize r) + n"
 
 fun 
   erase :: "arexp \<Rightarrow> rexp"
@@ -116,6 +146,7 @@
 | "erase (AALTs bs (r#rs)) = ALT (erase r) (erase (AALTs bs rs))"
 | "erase (ASEQ _ r1 r2) = SEQ (erase r1) (erase r2)"
 | "erase (ASTAR _ r) = STAR (erase r)"
+| "erase (ANTIMES _ r n) = NTIMES (erase r) n"
 
 
 fun fuse :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp" where
@@ -125,6 +156,7 @@
 | "fuse bs (AALTs cs rs) = AALTs (bs @ cs) rs"
 | "fuse bs (ASEQ cs r1 r2) = ASEQ (bs @ cs) r1 r2"
 | "fuse bs (ASTAR cs r) = ASTAR (bs @ cs) r"
+| "fuse bs (ANTIMES cs r n) = ANTIMES (bs @ cs) r n"
 
 lemma fuse_append:
   shows "fuse (bs1 @ bs2) r = fuse bs1 (fuse bs2 r)"
@@ -141,6 +173,7 @@
                                 (fuse [S]  (intern r2))"
 | "intern (SEQ r1 r2) = ASEQ [] (intern r1) (intern r2)"
 | "intern (STAR r) = ASTAR [] (intern r)"
+| "intern (NTIMES r n) = ANTIMES [] (intern r) n"
 
 
 fun retrieve :: "arexp \<Rightarrow> val \<Rightarrow> bit list" where
@@ -153,7 +186,9 @@
 | "retrieve (ASTAR bs r) (Stars []) = bs @ [S]"
 | "retrieve (ASTAR bs r) (Stars (v#vs)) = 
      bs @ [Z] @ retrieve r v @ retrieve (ASTAR [] r) (Stars vs)"
-
+| "retrieve (ANTIMES bs r 0) (Stars []) = bs @ [S]"
+| "retrieve (ANTIMES bs r (Suc n)) (Stars (v#vs)) = 
+     bs @ [Z] @ retrieve r v @ retrieve (ANTIMES [] r n) (Stars vs)"
 
 
 fun
@@ -165,27 +200,44 @@
 | "bnullable (AALTs bs rs) = (\<exists>r \<in> set rs. bnullable r)"
 | "bnullable (ASEQ bs r1 r2) = (bnullable r1 \<and> bnullable r2)"
 | "bnullable (ASTAR bs r) = True"
+| "bnullable (ANTIMES bs r n) = (if n  = 0 then True else bnullable r)"
 
 abbreviation
   bnullables :: "arexp list \<Rightarrow> bool"
 where
   "bnullables rs \<equiv> (\<exists>r \<in> set rs. bnullable r)"
 
-fun 
-  bmkeps :: "arexp \<Rightarrow> bit list" and
-  bmkepss :: "arexp list \<Rightarrow> bit list"
+function (sequential)
+  bmkeps :: "arexp \<Rightarrow> bit list" 
 where
   "bmkeps(AONE bs) = bs"
 | "bmkeps(ASEQ bs r1 r2) = bs @ (bmkeps r1) @ (bmkeps r2)"
-| "bmkeps(AALTs bs rs) = bs @ (bmkepss rs)"
+| "bmkeps(AALTs bs (r#rs)) = 
+    (if bnullable(r) then (bs @ bmkeps r) else (bmkeps (AALTs bs rs)))"
 | "bmkeps(ASTAR bs r) = bs @ [S]"
-| "bmkepss (r # rs) = (if bnullable(r) then (bmkeps r) else (bmkepss rs))"
+| "bmkeps(ANTIMES bs r 0) = bs @ [S]"
+| "bmkeps(ANTIMES bs r (Suc n)) = bs @ [Z] @ (bmkeps r) @ bmkeps(ANTIMES [] r n)"
+apply(pat_completeness)
+apply(auto)
+done
+
+termination "bmkeps"  
+apply(relation "measure asize") 
+        apply(auto)
+  using asize.elims by force
+
+fun
+   bmkepss :: "arexp list \<Rightarrow> bit list"
+where
+  "bmkepss (r # rs) = (if bnullable(r) then (bmkeps r) else (bmkepss rs))"
+
 
 lemma bmkepss1:
   assumes "\<not> bnullables rs1"
   shows "bmkepss (rs1 @ rs2) = bmkepss rs2"
   using assms
-  by (induct rs1) (auto)
+  by(induct rs1) (auto)
+  
 
 lemma bmkepss2:
   assumes "bnullables rs1"
@@ -206,7 +258,7 @@
       then AALT bs (ASEQ [] (bder c r1) r2) (fuse (bmkeps r1) (bder c r2))
       else ASEQ bs (bder c r1) r2)"
 | "bder c (ASTAR bs r) = ASEQ (bs @ [Z]) (bder c r) (ASTAR [] r)"
-
+| "bder c (ANTIMES bs r n) = (if n = 0 then AZERO else ASEQ (bs @ [Z]) (bder c r) (ANTIMES [] r (n - 1)))"
 
 fun 
   bders :: "arexp \<Rightarrow> string \<Rightarrow> arexp"
@@ -264,6 +316,15 @@
   apply(simp_all)
   done
 
+lemma retrieve_encode_NTIMES:
+  assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> code v = retrieve (intern r) v" "length vs = n"
+  shows "code (Stars vs) = retrieve (ANTIMES [] (intern r) n) (Stars vs)"
+  using assms
+  apply(induct vs arbitrary: n)
+   apply(simp_all)
+  by force
+
+
 lemma retrieve_fuse2:
   assumes "\<Turnstile> v : (erase r)"
   shows "retrieve (fuse bs r) v = bs @ retrieve r v"
@@ -285,7 +346,13 @@
   apply(auto elim!: Prf_elims)[1]
   apply(case_tac vs)
   apply(simp)
-  apply(simp)
+   apply(simp)
+  (* NTIMES *)
+  apply(auto elim!: Prf_elims)[1]
+  apply(case_tac vs1)
+   apply(simp_all)
+  apply(case_tac vs2)
+   apply(simp_all)
   done
 
 lemma retrieve_fuse:
@@ -300,8 +367,11 @@
   shows "code v = retrieve (intern r) v"
   using assms
   apply(induct v r )
-  apply(simp_all add: retrieve_fuse retrieve_encode_STARS)
-  done
+        apply(simp_all add: retrieve_fuse retrieve_encode_STARS)
+  apply(subst retrieve_encode_NTIMES)
+    apply(auto)
+  done 
+ 
 
 
 lemma retrieve_AALTs_bnullable1:
@@ -340,14 +410,26 @@
   apply (simp add: retrieve_AALTs_bnullable1)
   by (metis retrieve_AALTs_bnullable2)
 
-    
+lemma bmkeps_retrieve_ANTIMES: 
+  assumes "if n = 0 then True else bmkeps r = retrieve r (mkeps (erase r))" 
+  and     "bnullable (ANTIMES bs r n)"
+  shows "bmkeps (ANTIMES bs r n) = retrieve (ANTIMES bs r n) (Stars (replicate n (mkeps (erase r))))"
+ using assms
+  apply(induct n arbitrary: r bs)
+  apply(auto)[1]
+  apply(simp)
+  done
+
 lemma bmkeps_retrieve:
   assumes "bnullable r"
   shows "bmkeps r = retrieve r (mkeps (erase r))"
   using assms
-  apply(induct r)
-  apply(auto)  
-  using bmkeps_retrieve_AALTs by auto
+  apply(induct r rule: bmkeps.induct)
+        apply(auto)
+  apply (simp add: retrieve_AALTs_bnullable1)
+  using retrieve_AALTs_bnullable1 apply force
+  by (metis retrieve_AALTs_bnullable2)  
+  
 
 lemma bder_retrieve:
   assumes "\<Turnstile> v : der c (erase r)"
@@ -388,7 +470,15 @@
   apply(clarify)
   apply(erule Prf_elims)
   apply(clarify)
-  by (simp add: retrieve_fuse2)
+   apply (simp add: retrieve_fuse2)
+  (* ANTIMES case *)
+  apply(auto)  
+  apply(erule Prf_elims)
+  apply(erule Prf_elims)
+  apply(clarify)
+  apply(erule Prf_elims)
+  apply(clarify)
+  by (metis (full_types) Suc_pred append_assoc injval.simps(8) retrieve.simps(10) retrieve.simps(6))
 
 
 lemma MAIN_decode:
--- a/thys3/src/BlexerSimp.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/src/BlexerSimp.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -18,6 +18,7 @@
 | "(AALTs bs1 []) ~1 (AALTs bs2 []) = True"
 | "(AALTs bs1 (r1 # rs1)) ~1 (AALTs bs2 (r2 # rs2)) = (r1 ~1 r2 \<and> (AALTs bs1 rs1) ~1 (AALTs bs2 rs2))"
 | "(ASTAR bs1 r1) ~1 (ASTAR bs2 r2) = r1 ~1 r2"
+| "(ANTIMES bs1 r1 n1) ~1 (ANTIMES bs2 r2 n2) = (r1 ~1 r2 \<and> n1 = n2)"
 | "_ ~1 _ = False"
 
 
@@ -102,7 +103,12 @@
   assumes "bnullable r"
   shows "bmkeps (fuse bs r) = bs @ bmkeps r"
   using assms
-  by (induct r rule: bnullable.induct) (auto)
+  apply(induct r rule: bnullable.induct) 
+        apply(auto)
+  apply (metis append.assoc bmkeps_retrieve bmkeps_retrieve_AALTs bnullable.simps(4))  
+   apply(case_tac n)
+  apply(auto)
+  done
 
 lemma bmkepss_fuse: 
   assumes "bnullables rs"
@@ -363,7 +369,9 @@
   using rs1 srewrites7 apply presburger
   using srewrites7 apply force
   apply (simp add: srewrites7)
+   apply(simp add: srewrites7)
   by (simp add: srewrites7)
+  
 
 lemma bnullable0:
 shows "r1 \<leadsto> r2 \<Longrightarrow> bnullable r1 = bnullable r2" 
@@ -398,11 +406,21 @@
 next
   case (ss5 bs1 rs1 rsb)
   then show ?case
-    by (simp add: bmkepss1 bmkepss2 bmkepss_fuse bnullable_fuse)
+    apply (simp add: bmkepss1 bmkepss2 bmkepss_fuse bnullable_fuse)
+    apply(case_tac rs1)
+     apply(auto simp add: bnullable_fuse)
+    apply (metis bmkeps_retrieve bmkeps_retrieve_AALTs bnullable.simps(4))
+    apply (metis bmkeps_retrieve bmkeps_retrieve_AALTs bnullable.simps(4))
+    apply (metis bmkeps_retrieve bmkeps_retrieve_AALTs bnullable.simps(4))
+    by (metis bmkeps_retrieve bmkeps_retrieve_AALTs bnullable.simps(4))    
 next
   case (ss6 a1 a2 rsa rsb rsc)
   then show ?case
     by (smt (verit, best) Nil_is_append_conv bmkepss1 bmkepss2 bnullable_correctness in_set_conv_decomp list.distinct(1) list.set_intros(1) nullable_correctness set_ConsD subsetD)
+next
+  case (bs10 rs1 rs2 bs)
+  then show ?case
+    by (metis bmkeps_retrieve bmkeps_retrieve_AALTs bnullable.simps(4)) 
 qed (auto)
 
 lemma rewrites_bmkeps: 
--- a/thys3/src/ClosedForms.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/src/ClosedForms.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -539,8 +539,10 @@
      apply (simp add: frewrites_alt)
   apply (simp add: frewrites_cons)
    apply (simp add: frewrites_append)
-  by (simp add: frewrites_cons)
-
+  apply (simp add: frewrites_cons)
+  apply (auto simp add: frewrites_cons)
+  using frewrite.intros(1) many_steps_later by blast
+  
 
 lemma gstar0:
   shows "rsa @ (rdistinct rs (set rsa)) \<leadsto>g* rsa @ (rdistinct rs (insert RZERO (set rsa)))"
@@ -642,7 +644,8 @@
   apply(induct r)
   apply simp+
    apply (metis list.set_intros(1))
-  by blast
+  apply blast
+  by simp
   
 
 
@@ -1047,7 +1050,7 @@
     apply (meson flts_gstar greal_trans grewrites_ralts_rsimpalts gstar_rdistinct)
 
    apply (simp add: grewrites_ralts hrewrites_list)
-  by simp
+  by simp_all
 
 lemma interleave_aux1:
   shows " RALT (RSEQ RZERO r1) r h\<leadsto>*  r"
@@ -1121,7 +1124,7 @@
    apply(subgoal_tac "rder x (RALTS xa) h\<leadsto>* rder x (rsimp (RALTS xa))")
   using hrewrites_simpeq apply presburger
   using interleave_star1 simp_hrewrites apply presburger
-  by simp
+  by simp_all
   
 
 
@@ -1248,11 +1251,15 @@
   apply simp+
   
   using created_by_seq.cases apply blast
-  
-  apply (meson created_by_seq.cases rrexp.distinct(19) rrexp.distinct(21))
-  apply (metis created_by_seq.simps rder.simps(5))
-   apply (smt (verit, ccfv_threshold) created_by_seq.simps list.set_intros(1) list.simps(8) list.simps(9) rder.simps(4) rrexp.distinct(25) rrexp.inject(3))
-  using created_by_seq.intros(1) by force
+      apply(auto)
+  apply (meson created_by_seq.cases rrexp.distinct(23) rrexp.distinct(25))
+  using created_by_seq.simps apply blast
+  apply (meson created_by_seq.simps)
+  using created_by_seq.intros(1) apply blast
+  apply (metis (no_types, lifting) created_by_seq.simps k0a list.set_intros(1) list.simps(8) list.simps(9) rrexp.distinct(31))
+  apply (simp add: created_by_seq.intros(1))
+  using created_by_seq.simps apply blast
+  by (simp add: created_by_seq.intros(1))
 
 lemma createdbyseq_left_creatable:
   shows "created_by_seq (RALT r1 r2) \<Longrightarrow> created_by_seq r1"
@@ -1473,8 +1480,6 @@
 | "hElem r = [r]"
 
 
-
-
 lemma cbs_ders_cbs:
   shows "created_by_star r \<Longrightarrow> created_by_star (rder c r)"
   apply(induct r rule: created_by_star.induct)
@@ -1491,11 +1496,6 @@
   apply simp
   using cbs_ders_cbs by auto
 
-(*
-lemma created_by_star_cases:
-  shows "created_by_star r \<Longrightarrow> \<exists>ra rb. (r = RALT ra rb \<and> created_by_star ra \<and> created_by_star rb) \<or> r = RSEQ ra rb "
-  by (meson created_by_star.cases)
-*)
 
 
 lemma hfau_pushin: 
@@ -1549,6 +1549,8 @@
   apply(subst stupdates_append[symmetric])
   using stupdates_join_general by blast
 
+
+
 lemma starders_hfau_also1:
   shows "hflat_aux (rders (RSTAR r) (c # xs)) = map (\<lambda>s1. RSEQ (rders r s1) (RSTAR r)) (star_updates xs r [[c]])"
   using star_hfau_induct by force
@@ -1566,7 +1568,7 @@
   apply simp
   apply (metis (no_types, lifting) append_Cons append_eq_append_conv2 gmany_steps_later greal_trans grewrite.intros(2) grewrites_append self_append_conv)
   apply simp
-  by simp
+  by simp_all
   
 
 
@@ -1600,7 +1602,9 @@
   apply(subgoal_tac "rsimp (RALT a aa) = rsimp (RALTS (hflat_aux a @ hflat_aux aa))")
   using hflat_aux.simps(1) apply presburger
   apply simp
-  using cbs_hfau_rsimpeq1 by fastforce
+  using cbs_hfau_rsimpeq1 apply(fastforce)
+  by simp
+  
 
 lemma star_closed_form1:
   shows "rsimp (rders (RSTAR r0) (c#s)) = 
@@ -1644,8 +1648,11 @@
   using hrewrites_simpeq srewritescf_alt1 apply fastforce
   using star_closed_form6_hrewrites by blast
 
+
+
+
 lemma stupdate_nonempty:
-  shows "\<forall>s \<in> set  Ss. s \<noteq> [] \<Longrightarrow> \<forall>s \<in> set (star_update c r Ss). s \<noteq> []"
+  shows "\<forall>s \<in> set Ss. s \<noteq> [] \<Longrightarrow> \<forall>s \<in> set (star_update c r Ss). s \<noteq> []"
   apply(induct Ss)
   apply simp
   apply(case_tac "rnullable (rders r a)")
@@ -1671,12 +1678,518 @@
 lemma star_closed_form:
   shows "rders_simp (RSTAR r0) (c#s) = 
 rsimp ( RALTS ( (map (\<lambda>s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) ))"
-  apply(induct s)
+  apply(case_tac s)
    apply simp
    apply (metis idem_after_simp1 rsimp.simps(1) rsimp.simps(6) rsimp_idem)
   using star_closed_form4 star_closed_form5 star_closed_form6 star_closed_form8 by presburger
 
 
-unused_thms
+
+
+fun nupdate :: "char \<Rightarrow> rrexp \<Rightarrow>  (string * nat) option  list \<Rightarrow> (string * nat) option  list" where
+  "nupdate c r [] = []"
+| "nupdate c r (Some (s, Suc n) # Ss) = (if (rnullable (rders r s)) 
+                                          then Some (s@[c], Suc n) # Some ([c], n) # (nupdate c r Ss) 
+                                          else Some ((s@[c]), Suc n)  # (nupdate c r Ss) 
+                                        )"
+| "nupdate c r (Some (s, 0) # Ss) =  (if (rnullable (rders r s)) 
+                                        then Some (s@[c], 0) # None # (nupdate c r Ss) 
+                                        else Some ((s@[c]), 0)  # (nupdate c r Ss) 
+                                      ) "
+| "nupdate c r (None # Ss) = (None # nupdate c r Ss)"
+
+
+fun nupdates :: "char list \<Rightarrow> rrexp \<Rightarrow> (string * nat) option list \<Rightarrow> (string * nat) option list"
+  where
+  "nupdates [] r Ss = Ss"
+| "nupdates (c # cs) r Ss = nupdates cs r (nupdate c r Ss)"
+
+fun ntset :: "rrexp \<Rightarrow> nat \<Rightarrow> string \<Rightarrow> (string * nat) option list" where
+  "ntset r (Suc n)  (c # cs) = nupdates cs r [Some ([c], n)]"
+| "ntset r 0 _ = [None]"
+| "ntset r _ [] = []"
+
+inductive created_by_ntimes :: "rrexp \<Rightarrow> bool" where
+  "created_by_ntimes RZERO"
+| "created_by_ntimes (RSEQ ra (RNTIMES rb n))"
+| "\<lbrakk>created_by_ntimes r1; created_by_ntimes r2\<rbrakk> \<Longrightarrow> created_by_ntimes (RALT r1 r2)"
+| "\<lbrakk>created_by_ntimes r \<rbrakk> \<Longrightarrow> created_by_ntimes (RALT r RZERO)"
+
+fun highest_power_aux :: "(string * nat) option list \<Rightarrow> nat \<Rightarrow> nat" where
+  "highest_power_aux [] n = n"
+| "highest_power_aux (None # rs) n = highest_power_aux rs n"
+| "highest_power_aux (Some (s, n) # rs) m = highest_power_aux rs (max n m)"
+
+fun hpower :: "(string * nat) option list \<Rightarrow> nat" where
+  "hpower rs =  highest_power_aux rs 0"
+                        
+
+lemma nupdate_mono:
+  shows " (highest_power_aux (nupdate c r optlist) m) \<le> (highest_power_aux optlist m)"
+  apply(induct optlist arbitrary: m)
+   apply simp
+  apply(case_tac a)
+   apply simp
+  apply(case_tac aa)
+  apply(case_tac b)
+   apply simp+
+  done
+
+lemma nupdate_mono1:
+  shows "hpower (nupdate c r optlist) \<le> hpower optlist"
+  by (simp add: nupdate_mono)
+
+
+
+lemma cbn_ders_cbn:
+  shows "created_by_ntimes r \<Longrightarrow> created_by_ntimes (rder c r)"
+  apply(induct r rule: created_by_ntimes.induct)
+    apply simp
+
+  using created_by_ntimes.intros(1) created_by_ntimes.intros(2) created_by_ntimes.intros(3) apply presburger
+  
+  apply (metis created_by_ntimes.simps rder.simps(5) rder.simps(7))
+  using created_by_star.intros(1) created_by_star.intros(2) apply auto[1]
+  using created_by_ntimes.intros(1) created_by_ntimes.intros(3) apply auto[1]
+  by (metis (mono_tags, lifting) created_by_ntimes.simps list.simps(8) list.simps(9) rder.simps(1) rder.simps(4))
+
+lemma ntimes_ders_cbn:
+  shows "created_by_ntimes (rders (RSEQ r' (RNTIMES r n)) s)"
+  apply(induct s rule: rev_induct)
+   apply simp
+  apply (simp add: created_by_ntimes.intros(2))
+  apply(subst rders_append)
+  using cbn_ders_cbn by auto
+
+lemma always0:
+  shows "rders RZERO s = RZERO"
+  apply(induct s)
+  by simp+
+
+lemma ntimes_ders_cbn1:
+  shows "created_by_ntimes (rders (RNTIMES r n) (c#s))"
+  apply(case_tac n)
+   apply simp
+  using always0 created_by_ntimes.intros(1) apply auto[1]
+  by (simp add: ntimes_ders_cbn)
+
+
+lemma ntimes_hfau_pushin: 
+  shows "created_by_ntimes r \<Longrightarrow> hflat_aux (rder c r) = concat (map hflat_aux (map (rder c) (hflat_aux r)))"
+  apply(induct r rule: created_by_ntimes.induct)
+  apply simp+
+  done
+
+
+abbreviation
+  "opterm r SN \<equiv>     case SN of
+                                Some (s, n) \<Rightarrow> RSEQ (rders r s) (RNTIMES r n)
+                            |   None \<Rightarrow> RZERO
+                     
+              
+"
+
+fun nonempty_string :: "(string * nat) option \<Rightarrow> bool" where
+  "nonempty_string None = True"
+| "nonempty_string (Some ([], n)) = False"
+| "nonempty_string (Some (c#s, n)) = True"
+
+
+lemma nupdate_nonempty:
+  shows "\<lbrakk>\<forall>opt \<in> set Ss. nonempty_string opt \<rbrakk> \<Longrightarrow> \<forall>opt \<in> set (nupdate c r Ss). nonempty_string opt"
+  apply(induct c r Ss rule: nupdate.induct)
+     apply(auto)
+  apply (metis Nil_is_append_conv neq_Nil_conv nonempty_string.simps(3))
+  by (metis Nil_is_append_conv neq_Nil_conv nonempty_string.simps(3))
+
+
+
+lemma nupdates_nonempty:
+  shows "\<lbrakk>\<forall>opt \<in> set Ss. nonempty_string opt \<rbrakk> \<Longrightarrow> \<forall>opt \<in> set (nupdates s r Ss). nonempty_string opt"
+  apply(induct s arbitrary: Ss)
+   apply simp
+  apply simp
+  using nupdate_nonempty by presburger
+
+lemma nullability1: shows "rnullable (rders r s) = rnullable (rders_simp r s)"
+  by (metis der_simp_nullability rders.simps(1) rders_simp.simps(1) rders_simp_same_simpders)
+
+lemma nupdate_induct1:
+  shows 
+  "concat (map (hflat_aux \<circ> (rder c \<circ> (opterm r)))  sl ) = 
+   map (opterm r) (nupdate c r sl)"
+  apply(induct sl)
+   apply simp
+  apply(simp add: rders_append)
+  apply(case_tac "a")
+   apply simp+
+  apply(case_tac "aa")
+  apply(case_tac "b")
+  apply(case_tac "rnullable (rders r ab)")
+  apply(subgoal_tac "rnullable (rders_simp r ab)")
+    apply simp
+  using rders.simps(1) rders.simps(2) rders_append apply presburger
+  using nullability1 apply blast
+   apply simp
+  using rders.simps(1) rders.simps(2) rders_append apply presburger
+  apply simp
+  using rders.simps(1) rders.simps(2) rders_append by presburger
+
+
+lemma nupdates_join_general:
+  shows  "concat (map hflat_aux (map (rder x) (map (opterm r) (nupdates xs r Ss))  )) =
+           map (opterm r) (nupdates (xs @ [x]) r Ss)"
+  apply(induct xs arbitrary: Ss)
+   apply (simp)
+  prefer 2
+   apply auto[1]
+  using nupdate_induct1 by blast
+
+
+lemma nupdates_join_general1:
+  shows  "concat (map (hflat_aux \<circ> (rder x) \<circ> (opterm r)) (nupdates xs r Ss)) =
+           map (opterm r) (nupdates (xs @ [x]) r Ss)"
+  by (metis list.map_comp nupdates_join_general)
+
+lemma nupdates_append: shows 
+"nupdates (s @ [c]) r Ss = nupdate c r (nupdates s r Ss)"
+  apply(induct s arbitrary: Ss)
+   apply simp
+  apply simp
+  done
+
+lemma nupdates_mono:
+  shows "highest_power_aux (nupdates s r optlist) m \<le> highest_power_aux optlist m"
+  apply(induct s rule: rev_induct)
+   apply simp
+  apply(subst nupdates_append)
+  by (meson le_trans nupdate_mono)
+
+lemma nupdates_mono1:
+  shows "hpower (nupdates s r optlist) \<le> hpower optlist"
+  by (simp add: nupdates_mono)
+
+
+(*"\<forall>r \<in> set (nupdates s r [Some ([c], n)]). r = None \<or>( \<exists>s' m. r = Some (s', m) \<and> m \<le> n)"*)
+lemma nupdates_mono2:
+  shows "hpower (nupdates s r [Some ([c], n)]) \<le> n"
+  by (metis highest_power_aux.simps(1) highest_power_aux.simps(3) hpower.simps max_nat.right_neutral nupdates_mono1)
+
+lemma hpow_arg_mono:
+  shows "m \<ge> n \<Longrightarrow> highest_power_aux rs m \<ge> highest_power_aux rs n"
+  apply(induct rs arbitrary: m n)
+   apply simp
+  apply(case_tac a)
+   apply simp
+  apply(case_tac aa)
+  apply simp
+  done
+
+
+lemma hpow_increase:
+  shows "highest_power_aux (a # rs') m \<ge> highest_power_aux rs' m"
+  apply(case_tac a)
+   apply simp
+  apply simp
+  apply(case_tac aa)
+  apply(case_tac b)
+   apply simp+
+  apply(case_tac "Suc nat > m")
+  using hpow_arg_mono max.cobounded2 apply blast
+  using hpow_arg_mono max.cobounded2 by blast
+
+lemma hpow_append:
+  shows "highest_power_aux (rsa @ rsb) m  = highest_power_aux rsb (highest_power_aux rsa m)"
+  apply (induct rsa arbitrary: rsb m)
+   apply simp
+  apply simp
+  apply(case_tac a)
+   apply simp
+  apply(case_tac aa)
+  apply simp
+  done
+
+lemma hpow_aux_mono:
+  shows "highest_power_aux (rsa @ rsb) m \<ge> highest_power_aux rsb m"
+  apply(induct rsa arbitrary: rsb rule: rev_induct)
+  apply simp
+  apply simp
+  using hpow_increase order.trans by blast
+ 
+
+
+
+lemma hpow_mono:
+  shows "hpower (rsa @ rsb) \<le> n \<Longrightarrow> hpower rsb \<le> n"
+  apply(induct rsb arbitrary: rsa)
+   apply simp
+  apply(subgoal_tac "hpower rsb \<le> n")
+  apply simp
+  apply (metis dual_order.trans hpow_aux_mono)
+  by (metis hpow_append hpow_increase hpower.simps nat_le_iff_add trans_le_add1)
+
+
+lemma hpower_rs_elems_aux:
+  shows "highest_power_aux rs k \<le> n \<Longrightarrow> \<forall>r\<in>set rs. r = None \<or> (\<exists>s' m. r = Some (s', m) \<and> m \<le> n)"
+apply(induct rs k arbitrary: n rule: highest_power_aux.induct)
+    apply(auto)
+  by (metis dual_order.trans highest_power_aux.simps(1) hpow_append hpow_aux_mono linorder_le_cases max.absorb1 max.absorb2)
+
 
+lemma hpower_rs_elems:
+  shows "hpower rs \<le> n \<Longrightarrow> \<forall>r \<in> set rs. r = None \<or>( \<exists>s' m. r = Some (s', m) \<and> m \<le> n)"
+  by (simp add: hpower_rs_elems_aux)
+
+lemma nupdates_elems_leqn:
+  shows "\<forall>r \<in> set (nupdates s r [Some ([c], n)]). r = None \<or>( \<exists>s' m. r = Some (s', m) \<and> m \<le> n)"
+  by (meson hpower_rs_elems nupdates_mono2)
+
+lemma ntimes_hfau_induct:
+  shows "hflat_aux (rders (RSEQ (rder c r) (RNTIMES r n)) s) =   
+      map (opterm r) (nupdates s r [Some ([c], n)])"
+  apply(induct s rule: rev_induct)
+   apply simp
+  apply(subst rders_append)+
+  apply simp
+  apply(subst nupdates_append)
+  apply(subgoal_tac "created_by_ntimes (rders (RSEQ (rder c r) (RNTIMES r n)) xs)")
+  prefer 2
+  apply (simp add: ntimes_ders_cbn)
+  apply(subst ntimes_hfau_pushin)
+   apply simp
+  apply(subgoal_tac "concat (map hflat_aux (map (rder x) (hflat_aux (rders (RSEQ (rder c r) (RNTIMES r n)) xs)))) =
+                     concat (map hflat_aux (map (rder x) ( map (opterm r) (nupdates xs r [Some ([c], n)])))) ")
+   apply(simp only:)
+  prefer 2
+   apply presburger
+  apply(subst nupdates_append[symmetric])  
+  using nupdates_join_general by blast
+
+
+(*nupdates s r [Some ([c], n)]*)
+lemma ntimes_ders_hfau_also1:
+  shows "hflat_aux (rders (RNTIMES r (Suc n)) (c # xs)) = map (opterm r) (nupdates xs r [Some ([c], n)])"
+  using ntimes_hfau_induct by force
+
+
+
+lemma hfau_rsimpeq2_ntimes:
+  shows "created_by_ntimes r \<Longrightarrow> rsimp r = rsimp ( (RALTS (hflat_aux r)))"
+  apply(induct r)
+       apply simp+
+  
+    apply (metis rsimp_seq_equal1)
+  prefer 2
+   apply simp
+  apply(case_tac x)
+   apply simp
+  apply(case_tac "list")
+   apply simp
+  
+  apply (metis idem_after_simp1)
+  apply(case_tac "lista")
+  prefer 2
+   apply (metis hflat_aux.simps(8) idem_after_simp1 list.simps(8) list.simps(9) rsimp.simps(2))
+  apply(subgoal_tac "rsimp (RALT a aa) = rsimp (RALTS (hflat_aux (RALT a aa)))")
+  apply simp
+  apply(subgoal_tac "rsimp (RALT a aa) = rsimp (RALTS (hflat_aux a @ hflat_aux aa))")
+  using hflat_aux.simps(1) apply presburger
+  apply simp
+  using cbs_hfau_rsimpeq1 apply(fastforce)
+  by simp
+  
+
+lemma ntimes_closed_form1:
+  shows "rsimp (rders (RNTIMES r (Suc n)) (c#s)) = 
+rsimp ( ( RALTS (  map (opterm r) (nupdates s r [Some ([c], n)]) )))"
+  apply(subgoal_tac "created_by_ntimes (rders (RNTIMES r (Suc n)) (c#s))")
+   apply(subst hfau_rsimpeq2_ntimes)
+  apply linarith
+  using ntimes_ders_hfau_also1 apply auto[1]
+  using ntimes_ders_cbn1 by blast
+
+
+lemma ntimes_closed_form2:
+  shows  "rsimp (rders_simp (RNTIMES r (Suc n)) (c#s) ) = 
+rsimp ( ( RALTS ( (map (opterm r ) (nupdates s r [Some ([c], n)]) ) )))"
+  by (metis list.distinct(1) ntimes_closed_form1 rders_simp_same_simpders rsimp_idem)
+
+
+lemma ntimes_closed_form3:
+  shows  "rsimp (rders_simp (RNTIMES r n) (c#s)) =   (rders_simp (RNTIMES r n) (c#s))"
+  by (metis list.distinct(1) rders_simp_same_simpders rsimp_idem)
+
+
+lemma ntimes_closed_form4:
+  shows " (rders_simp (RNTIMES r (Suc n)) (c#s)) = 
+rsimp ( ( RALTS ( (map (opterm r ) (nupdates s r [Some ([c], n)]) )  )))"
+  using ntimes_closed_form2 ntimes_closed_form3 
+  by metis
+
+
+
+
+lemma ntimes_closed_form5:
+  shows " rsimp (  RALTS (map (\<lambda>s1. RSEQ (rders r0 s1) (RNTIMES r n) )         Ss)) = 
+          rsimp (  RALTS (map (\<lambda>s1. rsimp (RSEQ (rders r0 s1) (RNTIMES r n)) ) Ss))"
+  by (smt (verit, ccfv_SIG) list.map_comp map_eq_conv o_apply simp_flatten_aux0)
+
+
+
+lemma ntimes_closed_form6_hrewrites:
+  shows "  
+(map (\<lambda>s1. (RSEQ (rsimp (rders r0 s1)) (RNTIMES r0 n)) ) Ss )
+ scf\<leadsto>*
+(map (\<lambda>s1. rsimp (RSEQ (rders r0 s1) (RNTIMES r0 n)) ) Ss )"
+  apply(induct Ss)
+  apply simp
+  apply (simp add: ss1)
+  by (metis (no_types, lifting) list.simps(9) rsimp.simps(1) rsimp_idem simp_hrewrites ss2)
+
+
+
+lemma ntimes_closed_form6:
+  shows " rsimp ( ( RALTS ( (map (\<lambda>s1. rsimp (RSEQ (rders r0 s1) (RNTIMES r0 n)) ) Ss )))) = 
+          rsimp ( ( RALTS ( (map (\<lambda>s1.  (RSEQ (rsimp (rders r0 s1)) (RNTIMES r0 n)) ) Ss ))))"
+  apply(subgoal_tac " map (\<lambda>s1.  (RSEQ (rsimp (rders r0 s1)) (RNTIMES r0 n)) ) Ss  scf\<leadsto>*
+                      map (\<lambda>s1.  rsimp (RSEQ  (rders r0 s1) (RNTIMES r0 n)) ) Ss ")
+  using hrewrites_simpeq srewritescf_alt1 apply fastforce
+  using ntimes_closed_form6_hrewrites by blast
+
+abbreviation
+  "optermsimp r SN \<equiv>     case SN of
+                                Some (s, n) \<Rightarrow> RSEQ (rders_simp r s) (RNTIMES r n)
+                            |   None \<Rightarrow> RZERO
+                     
+              
+"
+
+abbreviation
+  "optermOsimp r SN \<equiv>     case SN of
+                                Some (s, n) \<Rightarrow> rsimp (RSEQ (rders r s) (RNTIMES r n))
+                            |   None \<Rightarrow> RZERO
+                     
+              
+"
+
+abbreviation
+  "optermosimp r SN \<equiv> case SN of
+                              Some (s, n) \<Rightarrow> RSEQ (rsimp (rders r s)) (RNTIMES r n)
+                            | None \<Rightarrow> RZERO
+"
+
+lemma ntimes_closed_form51:
+  shows "rsimp (RALTS (map (opterm r) (nupdates s r [Some ([c], n)]))) =
+         rsimp (RALTS (map (rsimp \<circ> (opterm r)) (nupdates s r [Some ([c], n)])))"
+  by (metis map_map simp_flatten_aux0)
+
+
+
+lemma osimp_Osimp:
+  shows " nonempty_string sn \<Longrightarrow> optermosimp r sn = optermsimp r sn"
+  apply(induct rule: nonempty_string.induct)
+  apply force
+   apply auto[1]
+  apply simp
+  by (metis list.distinct(1) rders.simps(2) rders_simp.simps(2) rders_simp_same_simpders)
+
+
+
+lemma osimp_Osimp_list:
+  shows "\<forall>sn \<in> set snlist. nonempty_string sn \<Longrightarrow> map (optermosimp r) snlist = map (optermsimp r) snlist"
+  by (simp add: osimp_Osimp)
+
+
+lemma ntimes_closed_form8:
+  shows  
+"rsimp (RALTS (map (optermosimp r) (nupdates s r [Some ([c], n)]))) =
+ rsimp (RALTS (map (optermsimp r) (nupdates s r [Some ([c], n)])))"
+  apply(subgoal_tac "\<forall>opt \<in> set (nupdates s r [Some ([c], n)]). nonempty_string opt")
+  using osimp_Osimp_list apply presburger
+  by (metis list.distinct(1) list.set_cases nonempty_string.simps(3) nupdates_nonempty set_ConsD)
+
+
+
+lemma ntimes_closed_form9aux:
+  shows "\<forall>snopt \<in> set (nupdates s r [Some ([c], n)]). nonempty_string snopt"
+  by (metis list.distinct(1) list.set_cases nonempty_string.simps(3) nupdates_nonempty set_ConsD)
+
+lemma ntimes_closed_form9aux1:
+  shows  "\<forall>snopt \<in> set snlist. nonempty_string snopt \<Longrightarrow> 
+rsimp (RALTS (map (optermosimp r) snlist)) =
+rsimp (RALTS (map (optermOsimp r) snlist))"
+  apply(induct snlist)
+   apply simp+
+  apply(case_tac "a")
+   apply simp+
+  by (smt (z3) case_prod_conv idem_after_simp1 map_eq_conv nonempty_string.elims(2) o_apply option.simps(4) option.simps(5) rsimp.simps(1) rsimp.simps(7) rsimp_idem)
+  
+
+
+
+lemma ntimes_closed_form9:
+  shows  
+"rsimp (RALTS (map (optermosimp r) (nupdates s r [Some ([c], n)]))) =
+ rsimp (RALTS (map (optermOsimp r) (nupdates s r [Some ([c], n)])))"
+  using ntimes_closed_form9aux ntimes_closed_form9aux1 by presburger
+
+
+lemma ntimes_closed_form10rewrites_aux:
+  shows "  map (rsimp \<circ> (opterm r)) optlist scf\<leadsto>* 
+           map (optermOsimp r)      optlist"
+  apply(induct optlist)
+   apply simp
+   apply (simp add: ss1)
+  apply simp
+  apply(case_tac a)
+  using ss2 apply fastforce
+  using ss2 by force
+  
+
+lemma ntimes_closed_form10rewrites:
+  shows "  map (rsimp \<circ> (opterm r)) (nupdates s r [Some ([c], n)]) scf\<leadsto>* 
+           map (optermOsimp r) (nupdates s r [Some ([c], n)])"
+  using ntimes_closed_form10rewrites_aux by blast
+
+lemma ntimes_closed_form10:
+  shows "rsimp (RALTS (map (rsimp \<circ> (opterm r)) (nupdates s r [Some ([c], n)]))) = 
+         rsimp (RALTS (map (optermOsimp r) (nupdates s r [Some ([c], n)])))"
+  by (smt (verit, best) case_prod_conv hpower_rs_elems map_eq_conv nupdates_mono2 o_apply option.case(2) option.simps(4) rsimp.simps(3))
+
+
+lemma rders_simp_cons:
+  shows "rders_simp r (c # s) = rders_simp (rsimp (rder c r)) s"
+  by simp
+
+lemma rder_ntimes:
+  shows "rder c (RNTIMES r (Suc n)) = RSEQ (rder c r) (RNTIMES r n)"
+  by simp
+
+
+lemma ntimes_closed_form:
+  shows "rders_simp (RNTIMES r0 (Suc n)) (c#s) = 
+rsimp ( RALTS ( (map (optermsimp r0 ) (nupdates s r0 [Some ([c], n)]) ) ))"
+  apply (subst rders_simp_cons)
+  apply(subst rder_ntimes)  
+  using ntimes_closed_form10 ntimes_closed_form4 ntimes_closed_form51 ntimes_closed_form8 ntimes_closed_form9 by force
+
+
+
+
+
+
+(*
+lemma ntimes_closed_form:
+  assumes "s \<noteq> []"
+  shows "rders_simp (RNTIMES r (Suc n)) s = 
+rsimp ( RALTS  (     map 
+                     (\<lambda> optSN. case optSN of
+                                Some (s, n) \<Rightarrow> RSEQ (rders_simp r s) (RNTIMES r n)
+                            |   None \<Rightarrow> RZERO
+                     ) 
+                     (ntset r n s) 
+               )
+      )"
+  
+*)
 end
\ No newline at end of file
--- a/thys3/src/ClosedFormsBounds.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/src/ClosedFormsBounds.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -141,7 +141,7 @@
              apply(simp only:)
   apply(subgoal_tac "rsizes (rdistinct rs (insert RZERO (insert RONE noalts_set \<union> alts_set)))
                    \<le>  rsizes (RONE # rdistinct rs (insert RZERO (insert RONE noalts_set \<union> alts_set)))")
-  apply (smt (verit, best) dual_order.trans insert_iff rrexp.distinct(15))
+  apply (smt (verit, ccfv_threshold) dual_order.trans insertE rrexp.distinct(17))
   apply (metis (no_types, opaque_lifting)  le_add_same_cancel2 list.simps(9) sum_list.Cons zero_le)
             apply fastforce
            apply fastforce
@@ -151,8 +151,8 @@
   using rflts.simps(4) apply presburger
       apply simp
       apply(subgoal_tac "insert RONE (noalts_set \<union> corr_set) = (insert RONE noalts_set) \<union> corr_set")
-  apply(simp only:)
-  apply (metis Un_insert_left insertE rrexp.distinct(15))
+        apply(simp only:)
+  apply (metis Un_insert_left insertE rrexp.distinct(17))
       apply fastforce
      apply(case_tac "a \<in> noalts_set")
       apply simp
@@ -174,14 +174,14 @@
   apply(subgoal_tac "(insert a (noalts_set \<union> alts_set)) = (insert a noalts_set) \<union> alts_set")
         apply(simp only:)
         apply(subgoal_tac "noalts_set \<union> corr_set = (insert a noalts_set) \<union> corr_set")
-  apply(simp only:)
-  apply (metis insertE rrexp.distinct(21))
+          apply(simp only:)
+  apply (metis insertE nonalt.simps(1) nonalt.simps(4))
         apply blast
   
   apply fastforce
   apply force
-     apply simp
-     apply (metis Un_insert_left insert_iff rrexp.distinct(21))
+      apply simp
+  apply (metis Un_insert_left insertE nonalt.simps(1) nonalt.simps(4))
     apply(case_tac "a \<in> noalts_set")
      apply simp
   apply(subgoal_tac "a \<notin> alts_set")
@@ -204,14 +204,13 @@
         apply(subgoal_tac "noalts_set \<union> corr_set = (insert a noalts_set) \<union> corr_set")
   apply(simp only:)
 
-
-  apply (metis insertE rrexp.distinct(25))
+         apply (metis insertE rrexp.distinct(31))
   apply blast
   apply fastforce
   apply force
      apply simp
   
-    apply (metis Un_insert_left insertE rrexp.distinct(25))
+    apply (metis Un_insert_left insertE rrexp.distinct(31))
 
   using Suc_le_mono flts_size_reduction_alts apply presburger
      apply(case_tac "a \<in> noalts_set")
@@ -234,16 +233,42 @@
   apply(subgoal_tac "(insert a (noalts_set \<union> alts_set)) = (insert a noalts_set) \<union> alts_set")
         apply(simp only:)
         apply(subgoal_tac "noalts_set \<union> corr_set = (insert a noalts_set) \<union> corr_set")
-  apply(simp only:)
-  apply (metis insertE rrexp.distinct(29))
+       apply(simp only:)
+  apply (metis insertE rrexp.distinct(37))
 
         apply blast
   
   apply fastforce
   apply force
      apply simp
-  apply (metis Un_insert_left insert_iff rrexp.distinct(29))
-  done
+   apply (metis Un_insert_left insert_iff rrexp.distinct(37))
+  apply(case_tac "a \<in> noalts_set")
+      apply simp
+  apply(subgoal_tac "a \<notin> alts_set")
+     prefer 2
+      apply blast
+  apply(case_tac "a \<in> corr_set")
+      apply(subgoal_tac "noalts_set \<union> corr_set = insert a ( noalts_set  \<union> corr_set)")
+  prefer 2
+  apply fastforce
+   apply(simp only:)
+   apply(subgoal_tac "rsizes (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set))) \<le>
+               rsizes (rdistinct (a # rs) (insert RZERO (noalts_set \<union> alts_set)))")
+
+       apply(subgoal_tac "rsizes (rdistinct (rflts (a # rs)) ((insert a noalts_set) \<union> corr_set)) \<le>
+          rsizes (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \<union> alts_set)))")
+  apply fastforce
+       apply simp
+  apply(subgoal_tac "(insert a (noalts_set \<union> alts_set)) = (insert a noalts_set) \<union> alts_set")
+        apply(simp only:)
+        apply(subgoal_tac "noalts_set \<union> corr_set = (insert a noalts_set) \<union> corr_set")
+       apply(simp only:)
+  apply (metis insertE nonalt.simps(1) nonalt.simps(7))
+  apply blast
+  apply blast
+  apply force
+  apply(auto)
+  by (metis Un_insert_left insert_iff rrexp.distinct(39))
 
 
 lemma flts_vs_nflts:
@@ -379,6 +404,220 @@
 qed
 
 
+thm ntimes_closed_form
+
+thm rsize.simps
+
+lemma nupdates_snoc:
+  shows " (nupdates (xs @ [x]) r optlist) = nupdate x r (nupdates xs r optlist)"
+  by (simp add: nupdates_append)
+
+lemma nupdate_elems:
+  shows "\<forall>opt \<in> set (nupdate c r optlist). opt = None \<or> (\<exists>s n. opt = Some (s, n))"
+  using nonempty_string.cases by auto
+
+lemma nupdates_elems:
+  shows "\<forall>opt \<in> set (nupdates s r optlist). opt = None \<or> (\<exists>s n. opt = Some (s, n))"
+  by (meson nonempty_string.cases)
+
+
+lemma opterm_optlist_result_shape:
+  shows "\<forall>r' \<in> set (map (optermsimp r) optlist). r' = RZERO \<or> (\<exists>s m. r' = RSEQ (rders_simp r s) (RNTIMES r m))"
+  apply(induct optlist)
+   apply simp
+  apply(case_tac a)
+  apply simp+
+  by fastforce
+
+
+lemma opterm_optlist_result_shape2:
+  shows "\<And>optlist. \<forall>r' \<in> set (map (optermsimp r) optlist). r' = RZERO \<or> (\<exists>s m. r' = RSEQ (rders_simp r s) (RNTIMES r m))"  
+  using opterm_optlist_result_shape by presburger
+
+
+lemma nupdate_n_leq_n:
+  shows "\<forall>r \<in> set (nupdate c' r [Some ([c], n)]). r = None \<or>( \<exists>s' m. r = Some (s', m) \<and> m \<le> n)"
+  apply(case_tac n)
+   apply simp
+  apply simp
+  done
+(*
+lemma nupdate_induct_leqn:
+  shows "\<lbrakk>\<forall>opt \<in> set optlist. opt = None \<or> (\<exists>s' m. opt = Some(s', m) \<and> m \<le> n) \<rbrakk> \<Longrightarrow> 
+       \<forall>opt \<in> set (nupdate c' r optlist). opt = None \<or> (\<exists>s' m. opt = Some (s', m) \<and> m \<le> n)"
+  apply (case_tac optlist)
+   apply simp
+  apply(case_tac a)
+   apply simp
+  sledgehammer
+*)
+
+
+lemma nupdates_n_leq_n:
+  shows "\<forall>r \<in> set (nupdates s r [Some ([c], n)]). r = None \<or>( \<exists>s' m. r = Some (s', m) \<and> m \<le> n)"
+  apply(induct s rule: rev_induct)
+   apply simp
+  apply(subst nupdates_append)
+  by (metis nupdates_elems_leqn nupdates_snoc)
+  
+
+
+lemma ntimes_closed_form_list_elem_shape:
+  shows "\<forall>r' \<in> set (map (optermsimp r) (nupdates s r [Some ([c], n)])). 
+r' = RZERO \<or> (\<exists>s' m. r' = RSEQ (rders_simp r s') (RNTIMES r m) \<and> m \<le> n)"
+  apply(insert opterm_optlist_result_shape2)
+  apply(case_tac s)
+   apply(auto)
+  apply (metis rders_simp_one_char)
+  by (metis case_prod_conv nupdates.simps(2) nupdates_n_leq_n option.simps(4) option.simps(5))
+
+
+lemma ntimes_trivial1:
+  shows "rsize RZERO \<le> N + rsize (RNTIMES r n)"
+  by simp
+
+
+lemma ntimes_trivial20:
+  shows "m \<le> n \<Longrightarrow> rsize (RNTIMES r m) \<le> rsize (RNTIMES r n)"
+  by simp
+
+
+lemma ntimes_trivial2:
+  assumes "\<forall>s. rsize (rders_simp r s) \<le> N"
+  shows "    r' = RSEQ (rders_simp r s1) (RNTIMES r m) \<and> m \<le> n
+       \<Longrightarrow> rsize r' \<le> Suc (N + rsize (RNTIMES r n))"
+  apply simp
+  by (simp add: add_mono_thms_linordered_semiring(1) assms)
+
+lemma ntimes_closed_form_list_elem_bounded:
+  assumes "\<forall>s. rsize (rders_simp r s) \<le> N"
+  shows "\<forall>r' \<in>  set  (map (optermsimp r) (nupdates s r [Some ([c], n)])). rsize r' \<le> Suc (N + rsize (RNTIMES r n))"
+  apply(rule ballI)
+  apply(subgoal_tac  "r' = RZERO \<or> (\<exists>s' m. r' = RSEQ (rders_simp r s') (RNTIMES r m) \<and> m \<le> n)")
+  prefer 2
+  using ntimes_closed_form_list_elem_shape apply blast
+  apply(case_tac "r' = RZERO")
+  using le_SucI ntimes_trivial1 apply presburger
+  apply(subgoal_tac "\<exists>s1 m. r' = RSEQ (rders_simp r s1) (RNTIMES r m) \<and> m \<le> n")
+  apply(erule exE)+
+  using assms ntimes_trivial2 apply presburger
+  by blast
+
+
+lemma P_holds_after_distinct:
+  assumes "\<forall>r \<in> set rs. P r"
+  shows "\<forall>r \<in> set (rdistinct rs rset). P r"
+  by (simp add: assms rdistinct_set_equality1)
+
+lemma ntimes_control_bounded:
+  assumes "\<forall>s. rsize (rders_simp r s) \<le> N"
+  shows "rsizes (rdistinct (map (optermsimp r) (nupdates s r [Some ([c], n)])) {}) 
+     \<le> (card (sizeNregex (Suc (N + rsize (RNTIMES r n))))) * (Suc (N + rsize (RNTIMES r n)))"
+  apply(subgoal_tac "\<forall>r' \<in> set (rdistinct (map (optermsimp r) (nupdates s r [Some ([c], n)])) {}).
+          rsize r' \<le> Suc (N + rsize (RNTIMES r n))")
+   apply (meson distinct_list_rexp_upto rdistinct_same_set)
+  apply(subgoal_tac "\<forall>r' \<in> set (map (optermsimp r) (nupdates s r [Some ([c], n)])). rsize r' \<le> Suc (N + rsize (RNTIMES r n))")
+   apply (simp add: rdistinct_set_equality)
+  by (metis assms nat_le_linear not_less_eq_eq ntimes_closed_form_list_elem_bounded)
+
+
+
+lemma ntimes_closed_form_bounded0:
+  assumes "\<forall>s. rsize (rders_simp r s) \<le> N"
+  shows " (rders_simp (RNTIMES r 0) s)  = RZERO \<or> (rders_simp (RNTIMES r 0) s)  = RNTIMES r 0
+           "
+  apply(induct s)
+   apply simp
+  by (metis always0 list.simps(3) rder.simps(7) rders.simps(2) rders_simp_same_simpders rsimp.simps(3))
+
+lemma ntimes_closed_form_bounded1:
+  assumes "\<forall>s. rsize (rders_simp r s) \<le> N"
+  shows " rsize (rders_simp (RNTIMES r 0) s) \<le> max (rsize  RZERO) (rsize (RNTIMES r 0))"
+  
+  by (metis assms max.cobounded1 max.cobounded2 ntimes_closed_form_bounded0)
+
+lemma self_smaller_than_bound:
+  shows "\<forall>s. rsize (rders_simp r s) \<le> N \<Longrightarrow> rsize r \<le> N"
+  apply(drule_tac x = "[]" in spec)
+  apply simp
+  done
+
+lemma ntimes_closed_form_bounded_nil_aux:
+  shows "max (rsize  RZERO) (rsize (RNTIMES r 0)) = 1 + rsize r"
+  by auto
+
+lemma ntimes_closed_form_bounded_nil:
+  assumes "\<forall>s. rsize (rders_simp r s) \<le> N"
+  shows " rsize (rders_simp (RNTIMES r 0) s) \<le> 1 + rsize r"
+  using assms ntimes_closed_form_bounded1 by auto
+
+lemma ntimes_ineq1:
+  shows "(rsize (RNTIMES r n)) \<ge> 1 + rsize r"
+  by simp
+
+lemma ntimes_ineq2:
+  shows "1 + rsize r \<le>  
+max ((Suc (card (sizeNregex (Suc (N + rsize (RNTIMES r n))))) * (Suc (N + rsize (RNTIMES r n))))) (rsize (RNTIMES r n))"
+  by (meson le_max_iff_disj ntimes_ineq1)
+
+lemma ntimes_closed_form_bounded:
+  assumes "\<forall>s. rsize (rders_simp r s) \<le> N"
+  shows "rsize (rders_simp (RNTIMES r (Suc n)) s) \<le> 
+           max ((Suc (card (sizeNregex (Suc (N + rsize (RNTIMES r n))))) * (Suc (N + rsize (RNTIMES r n))))) (rsize (RNTIMES r n))"
+proof(cases s)
+  case Nil
+  then show "rsize (rders_simp (RNTIMES r (Suc n)) s)
+    \<le> max (Suc (card (sizeNregex (Suc (N + rsize (RNTIMES r n))))) * Suc (N + rsize (RNTIMES r n))) (rsize (RNTIMES r n))" 
+    by simp
+next
+  case (Cons a list)
+
+  then have "rsize (rders_simp (RNTIMES r (Suc n)) s) = 
+             rsize (rsimp (RALTS ((map (optermsimp r)    (nupdates list r [Some ([a], n)])))))"
+    using ntimes_closed_form by fastforce
+  also have "... \<le> Suc (rsizes (rdistinct ((map (optermsimp r) (nupdates list r [Some ([a], n)]))) {}))"
+    using alts_simp_control by blast 
+  also have "... \<le> Suc (card (sizeNregex (Suc (N + rsize (RNTIMES r n))))) * (Suc (N + rsize (RNTIMES r n)))" 
+    using ntimes_control_bounded[OF assms]
+    by (metis add_mono le_add1 mult_Suc plus_1_eq_Suc)
+  also have "... \<le> max (Suc (card (sizeNregex (Suc (N + rsize (RNTIMES r n))))) * Suc (N + rsize (RNTIMES r n))) (rsize (RNTIMES r n))"
+    by simp    
+  finally show ?thesis by simp  
+qed
+
+
+lemma ntimes_closed_form_boundedA:
+  assumes "\<forall>s. rsize (rders_simp r s) \<le> N"
+  shows "\<exists>N'. \<forall>s. rsize (rders_simp (RNTIMES r n) s) \<le> N'"
+  apply(case_tac n)
+  using assms ntimes_closed_form_bounded_nil apply blast
+  using assms ntimes_closed_form_bounded by blast
+
+
+lemma star_closed_form_nonempty_bounded:
+  assumes "\<forall>s. rsize (rders_simp r s) \<le> N" and "s \<noteq> []"
+  shows "rsize (rders_simp (RSTAR r) s) \<le> 
+            ((Suc (card (sizeNregex (Suc (N + rsize (RSTAR r))))) * (Suc (N + rsize (RSTAR r))))) "
+proof(cases s)
+  case Nil
+  then show ?thesis 
+    using local.Nil by fastforce
+next
+  case (Cons a list)
+  then have "rsize (rders_simp (RSTAR r) s) = 
+    rsize (rsimp (RALTS ((map (\<lambda>s1. RSEQ (rders_simp r s1) (RSTAR r)) (star_updates list r [[a]])))))"
+    using star_closed_form by fastforce
+  also have "... \<le> Suc (rsizes (rdistinct (map (\<lambda>s1. RSEQ (rders_simp r s1) (RSTAR r)) (star_updates list r [[a]])) {}))"
+    using alts_simp_control by blast 
+  also have "... \<le> Suc (card (sizeNregex (Suc (N + rsize (RSTAR r))))) * (Suc (N + rsize (RSTAR r)))" 
+    by (smt (z3) add_mono_thms_linordered_semiring(1) assms(1) le_add1 map_eq_conv mult_Suc plus_1_eq_Suc star_control_bounded)
+  also have "... \<le> max (Suc (card (sizeNregex (Suc (N + rsize (RSTAR r))))) * Suc (N + rsize (RSTAR r))) (rsize (RSTAR r))"
+    by simp    
+  finally show ?thesis by simp  
+qed
+
+
+
 lemma seq_estimate_bounded: 
   assumes "\<forall>s. rsize (rders_simp r1 s) \<le> N1" 
       and "\<forall>s. rsize (rders_simp r2 s) \<le> N2"
@@ -425,7 +664,6 @@
     by auto 
 qed
 
-
 lemma rders_simp_bounded: 
   shows "\<exists>N. \<forall>s. rsize (rders_simp r s) \<le> N"
   apply(induct r)
@@ -440,9 +678,12 @@
   apply(assumption)
   apply(assumption)
   apply (metis alts_closed_form_bounded size_list_estimation')
-  using star_closed_form_bounded by blast
+  using star_closed_form_bounded apply blast
+  using ntimes_closed_form_boundedA by blast
+  
+  
+unused_thms
+export_code rders_simp rsimp rder in Scala module_name Example
 
 
-unused_thms
-
 end
--- a/thys3/src/FBound.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/src/FBound.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -18,6 +18,7 @@
 | "rerase (AALTs bs rs) = RALTS (map rerase rs)"
 | "rerase (ASEQ _ r1 r2) = RSEQ (rerase r1) (rerase r2)"
 | "rerase (ASTAR _ r) = RSTAR (rerase r)"
+| "rerase (ANTIMES _ r n) = RNTIMES (rerase r) n"
 
 lemma eq1_rerase:
   shows "x ~1 y \<longleftrightarrow> (rerase x) = (rerase y)"
--- a/thys3/src/GeneralRegexBound.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/src/GeneralRegexBound.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -18,6 +18,10 @@
 definition RALTs_set where
   "RALTs_set A n \<equiv> {RALTS rs | rs. \<forall>r \<in> set rs. r \<in> A \<and> rsizes rs \<le> n}"
 
+definition RNTIMES_set where
+  "RNTIMES_set A n \<equiv> {RNTIMES r m | m r. r \<in> A \<and> rsize r + m \<le> n}"
+
+
 definition
   "sizeNregex N \<equiv> {r. rsize r \<le> N}"
 
@@ -26,7 +30,8 @@
   "sizeNregex (Suc n) = (({RZERO, RONE} \<union> {RCHAR c| c. True}) 
                          \<union> (RSTAR ` sizeNregex n) 
                          \<union> (RSEQ_set (sizeNregex n) n)
-                         \<union> (RALTs_set (sizeNregex n) n))"
+                         \<union> (RALTs_set (sizeNregex n) n))
+                         \<union> (RNTIMES_set (sizeNregex n) n)"
   apply(auto)
         apply(case_tac x)
              apply(auto simp add: RSEQ_set_def)
@@ -37,15 +42,21 @@
          apply (simp add: RALTs_set_def)
   apply (metis imageI list.set_map member_le_sum_list order_trans)
   apply (simp add: sizeNregex_def)
-  apply (simp add: sizeNregex_def)
+        apply (simp add: sizeNregex_def)
+  apply (simp add: RNTIMES_set_def)
   apply (simp add: sizeNregex_def)
   using sizeNregex_def apply force
   apply (simp add: sizeNregex_def)
   apply (simp add: sizeNregex_def)
-  apply (simp add: RALTs_set_def)
+  apply (simp add: sizeNregex_def)
+    apply (simp add: RALTs_set_def)
   apply(simp add: sizeNregex_def)
   apply(auto)
-  using ex_in_conv by fastforce
+  using ex_in_conv apply fastforce
+  apply (simp add: RNTIMES_set_def)
+  apply(simp add: sizeNregex_def)
+  by force
+  
 
 lemma s4:
   "RSEQ_set A n \<subseteq> RSEQ_set_cartesian A"
@@ -155,6 +166,22 @@
    apply simp
   by (simp add: full_SetCompr_eq)
 
+thm RNTIMES_set_def
+
+lemma s9_aux0:
+  shows "RNTIMES_set (insert r A) n \<subseteq> RNTIMES_set A n \<union> (\<Union> i \<in> {..n}. {RNTIMES r i})"
+apply(auto simp add: RNTIMES_set_def)
+  done
+
+lemma s9_aux:
+  assumes "finite A"
+  shows "finite (RNTIMES_set A n)"
+  using assms
+  apply(induct A arbitrary: n)
+   apply(auto simp add: RNTIMES_set_def)[1]
+  apply(subgoal_tac "finite (RNTIMES_set F n \<union> (\<Union> i \<in> {..n}. {RNTIMES x i}))")
+  apply (metis finite_subset s9_aux0)
+  by blast
 
 lemma finite_size_n:
   shows "finite (sizeNregex n)"
@@ -175,8 +202,8 @@
   apply(rule finite_subset)
    apply(rule t2)
   apply(rule s8_aux)
-  apply(simp)
-  done
+   apply(simp)
+  by (simp add: s9_aux)
 
 lemma three_easy_cases0: 
   shows "rsize (rders_simp RZERO s) \<le> Suc 0"
--- a/thys3/src/Lexer.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/src/Lexer.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -3,7 +3,7 @@
   imports PosixSpec 
 begin
 
-section {* The Lexer Functions by Sulzmann and Lu  (without simplification) *}
+section \<open>The Lexer Functions by Sulzmann and Lu  (without simplification)\<close>
 
 fun 
   mkeps :: "rexp \<Rightarrow> val"
@@ -12,6 +12,7 @@
 | "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)"
 | "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))"
 | "mkeps(STAR r) = Stars []"
+| "mkeps(NTIMES r n) = Stars (replicate n (mkeps r))" 
 
 fun injval :: "rexp \<Rightarrow> char \<Rightarrow> val \<Rightarrow> val"
 where
@@ -22,6 +23,7 @@
 | "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2"
 | "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)"
 | "injval (STAR r) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" 
+| "injval (NTIMES r n) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" 
 
 fun 
   lexer :: "rexp \<Rightarrow> string \<Rightarrow> val option"
@@ -33,20 +35,32 @@
 
 
 
-section {* Mkeps, Injval Properties *}
+section \<open>Mkeps, Injval Properties\<close>
+
+lemma mkeps_flat:
+  assumes "nullable(r)" 
+  shows "flat (mkeps r) = []"
+using assms
+  by (induct rule: mkeps.induct) (auto)
+
+lemma Prf_NTimes_empty:
+  assumes "\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v = []" 
+  and     "length vs = n"
+  shows "\<Turnstile> Stars vs : NTIMES r n"
+  using assms
+  by (metis Prf.intros(7) empty_iff eq_Nil_appendI list.set(1))
+  
 
 lemma mkeps_nullable:
   assumes "nullable(r)" 
   shows "\<Turnstile> mkeps r : r"
 using assms
-by (induct rule: nullable.induct) 
-   (auto intro: Prf.intros)
-
-lemma mkeps_flat:
-  assumes "nullable(r)" 
-  shows "flat (mkeps r) = []"
-using assms
-by (induct rule: nullable.induct) (auto)
+  apply (induct rule: mkeps.induct) 
+  apply(auto intro: Prf.intros split: if_splits)
+  apply (metis Prf.intros(7) append_is_Nil_conv empty_iff list.set(1) list.size(3))
+  apply(rule Prf_NTimes_empty)
+  apply(auto simp add: mkeps_flat) 
+  done
 
 lemma Prf_injval_flat:
   assumes "\<Turnstile> v : der c r" 
@@ -62,14 +76,20 @@
 using assms
 apply(induct r arbitrary: c v rule: rexp.induct)
 apply(auto intro!: Prf.intros mkeps_nullable elim!: Prf_elims split: if_splits)
+(* Star *)
 apply(simp add: Prf_injval_flat)
-done
+(* NTimes *)
+  apply(case_tac x2)
+    apply(simp)
+  apply(simp)
+  apply(subst append.simps(2)[symmetric])
+  apply(rule Prf.intros)
+  apply(auto simp add: Prf_injval_flat)
+  done
 
 
+text \<open>Mkeps and injval produce, or preserve, Posix values.\<close>
 
-text {*
-  Mkeps and injval produce, or preserve, Posix values.
-*}
 
 lemma Posix_mkeps:
   assumes "nullable r"
@@ -80,7 +100,7 @@
 apply(subst append.simps(1)[symmetric])
 apply(rule Posix.intros)
 apply(auto)
-done
+by (simp add: Posix_NTIMES2 pow_empty_iff)
 
 lemma Posix_injval:
   assumes "s \<in> (der c r) \<rightarrow> v"
@@ -228,11 +248,60 @@
         ultimately 
         have "((c # s1) @ s2) \<in> STAR r \<rightarrow> Stars (injval r c v1 # vs)" by (rule Posix.intros)
         then show "(c # s) \<in> STAR r \<rightarrow> injval (STAR r) c v" using cons by(simp)
-    qed
+      qed
+next
+  case (NTIMES r n)
+  have IH: "\<And>s v. s \<in> der c r \<rightarrow> v \<Longrightarrow> (c # s) \<in> r \<rightarrow> injval r c v" by fact
+  have "s \<in> der c (NTIMES r n) \<rightarrow> v" by fact
+  then consider
+      (cons) v1 vs s1 s2 where 
+        "v = Seq v1 (Stars vs)" "s = s1 @ s2" 
+        "s1 \<in> der c r \<rightarrow> v1" "s2 \<in> (NTIMES r (n - 1)) \<rightarrow> (Stars vs)" "0 < n"
+        "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1)))" 
+    
+    apply(auto elim: Posix_elims simp add: der_correctness Der_def intro: Posix.intros split: if_splits)
+    apply(erule Posix_elims)
+    apply(simp)
+    apply(subgoal_tac "\<exists>vss. v2 = Stars vss")
+    apply(clarify)
+    apply(drule_tac x="vss" in meta_spec)
+    apply(drule_tac x="s1" in meta_spec)
+    apply(drule_tac x="s2" in meta_spec)
+     apply(simp add: der_correctness Der_def)
+    apply(erule Posix_elims)
+     apply(auto)
+      done
+    then show "(c # s) \<in> (NTIMES r n) \<rightarrow> injval (NTIMES r n) c v" 
+    proof (cases)
+      case cons
+          have "s1 \<in> der c r \<rightarrow> v1" by fact
+          then have "(c # s1) \<in> r \<rightarrow> injval r c v1" using IH by simp
+        moreover
+          have "s2 \<in> (NTIMES r (n - 1)) \<rightarrow> Stars vs" by fact
+        moreover 
+          have "(c # s1) \<in> r \<rightarrow> injval r c v1" by fact 
+          then have "flat (injval r c v1) = (c # s1)" by (rule Posix1)
+          then have "flat (injval r c v1) \<noteq> []" by simp
+        moreover 
+          have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L (der c r) \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1)))" by fact
+          then have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (c # s1) @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1)))"
+            by (simp add: der_correctness Der_def)
+        ultimately 
+        have "((c # s1) @ s2) \<in> NTIMES r n \<rightarrow> Stars (injval r c v1 # vs)" 
+           apply (rule_tac Posix.intros)
+               apply(simp_all)
+              apply(case_tac n)
+            apply(simp)
+           using Posix_elims(1) NTIMES.prems apply auto[1]
+             apply(simp)
+             done
+        then show "(c # s) \<in> NTIMES r n \<rightarrow> injval (NTIMES r n) c v" using cons by(simp)
+      qed  
+
 qed
 
 
-section {* Lexer Correctness *}
+section \<open>Lexer Correctness\<close>
 
 
 lemma lexer_correct_None:
@@ -354,7 +423,8 @@
      apply(erule Prf_elims)
   apply(auto)
    apply (smt Prf_elims(6) injval.simps(7) list.inject val.inject(5))
-  by (smt Prf_elims(6) injval.simps(7) list.inject val.inject(5))
+  apply (smt Prf_elims(6) injval.simps(7) list.inject val.inject(5))
+  by (smt (verit, best) Prf_elims(1) Prf_elims(2) Prf_elims(7) injval.simps(8) list.inject val.simps(5))
   
   
 
@@ -375,7 +445,7 @@
    apply (simp add: lexer_correctness(1))
   apply(subgoal_tac "\<Turnstile> a : (der c r)")
    prefer 2
-  using Posix_Prf apply blast
+  using Posix1a apply blast
   using injval_inj by blast
   
 
--- a/thys3/src/PosixSpec.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/src/PosixSpec.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -46,6 +46,9 @@
 | "\<Turnstile> Void : ONE"
 | "\<Turnstile> Char c : CH c"
 | "\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> [] \<Longrightarrow> \<Turnstile> Stars vs : STAR r"
+| "\<lbrakk>\<forall>v \<in> set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []; 
+    \<forall>v \<in> set vs2. \<Turnstile> v : r \<and> flat v = []; 
+    length (vs1 @ vs2) = n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (vs1 @ vs2) : NTIMES r n"   
 
 inductive_cases Prf_elims:
   "\<Turnstile> v : ZERO"
@@ -54,6 +57,7 @@
   "\<Turnstile> v : ONE"
   "\<Turnstile> v : CH c"
   "\<Turnstile> vs : STAR r"
+  "\<Turnstile> vs : NTIMES r n"
 
 lemma Prf_Stars_appendE:
   assumes "\<Turnstile> Stars (vs1 @ vs2) : STAR r"
@@ -61,6 +65,28 @@
 using assms
 by (auto intro: Prf.intros elim!: Prf_elims)
 
+lemma Pow_cstring:
+  fixes A::"string set"
+  assumes "s \<in> A ^^ n"
+  shows "\<exists>ss1 ss2. concat (ss1 @ ss2) = s \<and> length (ss1 @ ss2) = n \<and> 
+         (\<forall>s \<in> set ss1. s \<in> A \<and> s \<noteq> []) \<and> (\<forall>s \<in> set ss2. s \<in> A \<and> s = [])"
+using assms
+apply(induct n arbitrary: s)
+  apply(auto)[1]
+  apply(auto simp add: Sequ_def)
+  apply(drule_tac x="s2" in meta_spec)
+  apply(simp)
+apply(erule exE)+
+  apply(clarify)
+apply(case_tac "s1 = []")
+apply(simp)
+apply(rule_tac x="ss1" in exI)
+apply(rule_tac x="s1 # ss2" in exI)
+apply(simp)
+apply(rule_tac x="s1 # ss1" in exI)
+apply(rule_tac x="ss2" in exI)
+  apply(simp)
+  done
 
 lemma flats_Prf_value:
   assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r"
@@ -75,15 +101,48 @@
 apply(simp)
 apply(rule_tac x="v#vs" in exI)
 apply(simp)
-done
+  done
+
+lemma Aux:
+  assumes "\<forall>s\<in>set ss. s = []"
+  shows "concat ss = []"
+using assms
+by (induct ss) (auto)
 
+lemma flats_cval:
+  assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r"
+  shows "\<exists>vs1 vs2. flats (vs1 @ vs2) = concat ss \<and> length (vs1 @ vs2) = length ss \<and> 
+          (\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []) \<and>
+          (\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = [])"
+using assms
+apply(induct ss rule: rev_induct)
+apply(rule_tac x="[]" in exI)+
+apply(simp)
+apply(simp)
+apply(clarify)
+apply(case_tac "flat v = []")
+apply(rule_tac x="vs1" in exI)
+apply(rule_tac x="v#vs2" in exI)
+apply(simp)
+apply(rule_tac x="vs1 @ [v]" in exI)
+apply(rule_tac x="vs2" in exI)
+apply(simp)
+by (simp add: Aux)
+
+lemma pow_Prf:
+  assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<in> A"
+  shows "flats vs \<in> A ^^ (length vs)"
+  using assms
+  by (induct vs) (auto)
 
 lemma L_flat_Prf1:
   assumes "\<Turnstile> v : r" 
   shows "flat v \<in> L r"
-using assms
-by (induct) (auto simp add: Sequ_def Star_concat)
-
+  using assms
+  apply (induct v r rule: Prf.induct) 
+  apply(auto simp add: Sequ_def Star_concat lang_pow_add)
+  by (metis pow_Prf)
+  
 lemma L_flat_Prf2:
   assumes "s \<in> L r" 
   shows "\<exists>v. \<Turnstile> v : r \<and> flat v = s"
@@ -105,7 +164,30 @@
 next
   case (ALT r1 r2 s)
   then show "\<exists>v. \<Turnstile> v : ALT r1 r2 \<and> flat v = s"
-  unfolding L.simps by (fastforce intro: Prf.intros)
+    unfolding L.simps by (fastforce intro: Prf.intros)
+next
+  case (NTIMES r n)
+  have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact
+  have "s \<in> L (NTIMES r n)" by fact
+  then obtain ss1 ss2 where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = n" 
+    "\<forall>s \<in> set ss1. s \<in> L r \<and> s \<noteq> []" "\<forall>s \<in> set ss2. s \<in> L r \<and> s = []"
+  using Pow_cstring by force
+  then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = n" 
+      "\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []" "\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = []"
+    using IH flats_cval 
+  apply -
+  apply(drule_tac x="ss1 @ ss2" in meta_spec)
+  apply(drule_tac x="r" in meta_spec)
+  apply(drule meta_mp)
+  apply(simp)
+  apply (metis Un_iff)
+  apply(clarify)
+  apply(drule_tac x="vs1" in meta_spec)
+  apply(drule_tac x="vs2" in meta_spec)
+  apply(simp)
+  done
+  then show "\<exists>v. \<Turnstile> v : NTIMES r n \<and> flat v = s"
+  using Prf.intros(7) flat_Stars by blast
 qed (auto intro: Prf.intros)
 
 
@@ -130,9 +212,11 @@
   and   "LV ONE s = (if s = [] then {Void} else {})"
   and   "LV (CH c) s = (if s = [c] then {Char c} else {})"
   and   "LV (ALT r1 r2) s = Left ` LV r1 s \<union> Right ` LV r2 s"
+  and   "LV (NTIMES r 0) s = (if s = [] then {Stars []} else {})"
 unfolding LV_def
-by (auto intro: Prf.intros elim: Prf.cases)
-
+  apply (auto intro: Prf.intros elim: Prf.cases)
+  by (metis Prf.intros(7) append.right_neutral empty_iff list.set(1) list.size(3))
+  
 
 abbreviation
   "Prefixes s \<equiv> {s'. prefix s' s}"
@@ -174,6 +258,64 @@
   ultimately show "finite (Prefixes s)" by simp
 qed
 
+definition
+  "Stars_Append Vs1 Vs2 \<equiv> {Stars (vs1 @ vs2) | vs1 vs2. Stars vs1 \<in> Vs1 \<and> Stars vs2 \<in> Vs2}"
+
+lemma finite_Stars_Append:
+  assumes "finite Vs1" "finite Vs2"
+  shows "finite (Stars_Append Vs1 Vs2)"
+  using assms  
+proof -
+  define UVs1 where "UVs1 \<equiv> Stars -` Vs1"
+  define UVs2 where "UVs2 \<equiv> Stars -` Vs2"  
+  from assms have "finite UVs1" "finite UVs2"
+    unfolding UVs1_def UVs2_def
+    by(simp_all add: finite_vimageI inj_on_def) 
+  then have "finite ((\<lambda>(vs1, vs2). Stars (vs1 @ vs2)) ` (UVs1 \<times> UVs2))"
+    by simp
+  moreover 
+    have "Stars_Append Vs1 Vs2 = (\<lambda>(vs1, vs2). Stars (vs1 @ vs2)) ` (UVs1 \<times> UVs2)"
+    unfolding Stars_Append_def UVs1_def UVs2_def by auto    
+  ultimately show "finite (Stars_Append Vs1 Vs2)"   
+    by simp
+qed 
+
+lemma LV_NTIMES_subset:
+  "LV (NTIMES r n) s \<subseteq> Stars_Append (LV (STAR r) s) (\<Union>i\<le>n. LV (NTIMES r i) [])"
+apply(auto simp add: LV_def)
+apply(auto elim!: Prf_elims)
+  apply(auto simp add: Stars_Append_def)
+  apply(rule_tac x="vs1" in exI)
+  apply(rule_tac x="vs2" in exI)  
+  apply(auto)
+    using Prf.intros(6) apply(auto)
+      apply(rule_tac x="length vs2" in bexI)
+    thm Prf.intros
+      apply(subst append.simps(1)[symmetric])
+    apply(rule Prf.intros)
+      apply(auto)[1]
+      apply(auto)[1]
+     apply(simp)
+    apply(simp)
+    done
+
+lemma LV_NTIMES_Suc_empty:
+  shows "LV (NTIMES r (Suc n)) [] = 
+     (\<lambda>(v, vs). Stars (v#vs)) ` (LV r [] \<times> (Stars -` (LV (NTIMES r n) [])))"
+unfolding LV_def
+apply(auto elim!: Prf_elims simp add: image_def)
+apply(case_tac vs1)
+apply(auto)
+apply(case_tac vs2)
+apply(auto)
+apply(subst append.simps(1)[symmetric])
+apply(rule Prf.intros)
+apply(auto)
+apply(subst append.simps(1)[symmetric])
+apply(rule Prf.intros)
+apply(auto)
+  done 
+
 lemma LV_STAR_finite:
   assumes "\<forall>s. finite (LV r s)"
   shows "finite (LV (STAR r) s)"
@@ -215,7 +357,22 @@
   ultimately
   show "finite (LV (STAR r) s)" by (simp add: finite_subset)
 qed  
-    
+
+lemma finite_NTimes_empty:
+  assumes "\<And>s. finite (LV r s)" 
+  shows "finite (LV (NTIMES r n) [])"
+  using assms
+  apply(induct n)
+   apply(auto simp add: LV_simps)
+  apply(subst LV_NTIMES_Suc_empty)
+  apply(rule finite_imageI)
+  apply(rule finite_cartesian_product)
+  using assms apply simp 
+  apply(rule finite_vimageI)
+  apply(simp)
+  apply(simp add: inj_on_def)
+  done
+
 
 lemma LV_finite:
   shows "finite (LV r s)"
@@ -251,6 +408,15 @@
 next
   case (STAR r s)
   then show "finite (LV (STAR r) s)" by (simp add: LV_STAR_finite)
+next
+  case (NTIMES r n s)
+  have "\<And>s. finite (LV r s)" by fact
+  then have "finite (Stars_Append (LV (STAR r) s) (\<Union>i\<le>n. LV (NTIMES r i) []))" 
+    apply(rule_tac finite_Stars_Append)
+     apply (simp add: LV_STAR_finite)
+    using finite_NTimes_empty by blast
+  then show "finite (LV (NTIMES r n) s)"
+    by (metis LV_NTIMES_subset finite_subset)
 qed
 
 
@@ -271,6 +437,11 @@
     \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))\<rbrakk>
     \<Longrightarrow> (s1 @ s2) \<in> STAR r \<rightarrow> Stars (v # vs)"
 | Posix_STAR2: "[] \<in> STAR r \<rightarrow> Stars []"
+| Posix_NTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> NTIMES r (n - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < n;
+    \<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1)))\<rbrakk>
+    \<Longrightarrow> (s1 @ s2) \<in> NTIMES r n \<rightarrow> Stars (v # vs)"
+| Posix_NTIMES2: "\<lbrakk>\<forall>v \<in> set vs. [] \<in> r \<rightarrow> v; length vs = n\<rbrakk>
+    \<Longrightarrow> [] \<in> NTIMES r n \<rightarrow> Stars vs"  
 
 inductive_cases Posix_elims:
   "s \<in> ZERO \<rightarrow> v"
@@ -279,19 +450,47 @@
   "s \<in> ALT r1 r2 \<rightarrow> v"
   "s \<in> SEQ r1 r2 \<rightarrow> v"
   "s \<in> STAR r \<rightarrow> v"
+  "s \<in> NTIMES r n \<rightarrow> v"
 
 lemma Posix1:
   assumes "s \<in> r \<rightarrow> v"
   shows "s \<in> L r" "flat v = s"
 using assms
-  by(induct s r v rule: Posix.induct)
-    (auto simp add: Sequ_def)
+  apply(induct s r v rule: Posix.induct)
+  apply(auto simp add: pow_empty_iff)
+  apply (metis Suc_pred concI lang_pow.simps(2))
+  by (meson ex_in_conv set_empty)
+
+lemma Posix1a:
+  assumes "s \<in> r \<rightarrow> v"
+  shows "\<Turnstile> v : r"
+using assms
+  apply(induct s r v rule: Posix.induct)
+  apply(auto intro: Prf.intros)
+  apply (metis Prf.intros(6) Prf_elims(6) set_ConsD val.inject(5))
+  prefer 2
+  apply (metis Posix1(2) Prf.intros(7) append_Nil empty_iff list.set(1))
+  apply(erule Prf_elims)
+  apply(auto)
+  apply(subst append.simps(2)[symmetric])
+  apply(rule Prf.intros)
+  apply(auto)
+  done
 
 text \<open>
   For a give value and string, our Posix definition 
   determines a unique value.
 \<close>
 
+lemma List_eq_zipI:
+  assumes "list_all2 (\<lambda>v1 v2. v1 = v2) vs1 vs2" 
+  and "length vs1 = length vs2"
+  shows "vs1 = vs2"  
+ using assms
+  apply(induct vs1 vs2 rule: list_all2_induct)
+  apply(auto)
+  done 
+
 lemma Posix_determ:
   assumes "s \<in> r \<rightarrow> v1" "s \<in> r \<rightarrow> v2"
   shows "v1 = v2"
@@ -356,6 +555,29 @@
   case (Posix_STAR2 r v2)
   have "[] \<in> STAR r \<rightarrow> v2" by fact
   then show "Stars [] = v2" by cases (auto simp add: Posix1)
+next
+  case (Posix_NTIMES2 vs r n v2)
+  then show "Stars vs = v2"
+    apply(erule_tac Posix_elims)
+    apply(auto)
+    apply (simp add: Posix1(2))  
+    apply(rule List_eq_zipI)
+     apply(auto simp add: list_all2_iff)
+    by (meson in_set_zipE)
+next
+  case (Posix_NTIMES1 s1 r v s2 n vs)
+  have "(s1 @ s2) \<in> NTIMES r n \<rightarrow> v2" 
+       "s1 \<in> r \<rightarrow> v" "s2 \<in> NTIMES r (n - 1) \<rightarrow> Stars vs" "flat v \<noteq> []"
+       "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1 )))" by fact+
+  then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (NTIMES r (n - 1)) \<rightarrow> (Stars vs')"
+  apply(cases) apply (auto simp add: append_eq_append_conv2)
+    using Posix1(1) apply fastforce
+    apply (metis One_nat_def Posix1(1) Posix_NTIMES1.hyps(7) append.right_neutral append_self_conv2)
+  using Posix1(2) by blast
+  moreover
+  have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2"
+            "\<And>v2. s2 \<in> NTIMES r (n - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+
+  ultimately show "Stars (v # vs) = v2" by auto
 qed
 
 
@@ -368,13 +590,9 @@
   shows "v \<in> LV r s"
   using assms unfolding LV_def
   apply(induct rule: Posix.induct)
-  apply(auto simp add: intro!: Prf.intros elim!: Prf_elims)
-  done
+   apply(auto simp add: intro!: Prf.intros elim!: Prf_elims Posix1a)
+   apply (smt (verit, best) One_nat_def Posix1a Posix_NTIMES1 L.simps(7))
+  using Posix1a Posix_NTIMES2 by blast
 
-lemma Posix_Prf:
-  assumes "s \<in> r \<rightarrow> v"
-  shows "\<Turnstile> v : r"
-  using assms Posix_LV LV_def
-  by simp
 
 end
--- a/thys3/src/RegLangs.thy	Wed Jul 13 08:35:09 2022 +0100
+++ b/thys3/src/RegLangs.thy	Wed Jul 13 08:35:45 2022 +0100
@@ -21,6 +21,42 @@
   and   "{} ;; A = {}"
   by (simp_all add: Sequ_def)
 
+lemma concI[simp,intro]: "u : A \<Longrightarrow> v : B \<Longrightarrow> u@v : A ;; B"
+by (auto simp add: Sequ_def)
+
+lemma concE[elim]: 
+assumes "w \<in> A ;; B"
+obtains u v where "u \<in> A" "v \<in> B" "w = u@v"
+using assms by (auto simp: Sequ_def)
+
+lemma concI_if_Nil2: "[] \<in> B \<Longrightarrow> xs : A \<Longrightarrow> xs \<in> A ;; B"
+by (metis append_Nil2 concI)
+
+lemma conc_assoc: "(A ;; B) ;; C = A ;; (B ;; C)"
+by (auto elim!: concE) (simp only: append_assoc[symmetric] concI)
+
+
+text \<open>Language power operations\<close>
+
+overloading lang_pow == "compow :: nat \<Rightarrow> string set \<Rightarrow> string set"
+begin
+  primrec lang_pow :: "nat \<Rightarrow> string set \<Rightarrow> string set" where
+  "lang_pow 0 A = {[]}" |
+  "lang_pow (Suc n) A = A ;; (lang_pow n A)"
+end
+
+
+lemma conc_pow_comm:
+  shows "A ;; (A ^^ n) = (A ^^ n) ;; A"
+by (induct n) (simp_all add: conc_assoc[symmetric])
+
+lemma lang_pow_add: "A ^^ (n + m) = (A ^^ n) ;; (A ^^ m)"
+  by (induct n) (auto simp: conc_assoc)
+
+lemma lang_empty: 
+  fixes A::"string set"
+  shows "A ^^ 0 = {[]}"
+  by simp
 
 section \<open>Semantic Derivative (Left Quotient) of Languages\<close>
 
@@ -88,6 +124,11 @@
 unfolding Der_def Sequ_def
 by(auto simp add: Star_decomp)
 
+lemma Der_inter[simp]:   "Der a (A \<inter> B) = Der a A \<inter> Der a B"
+  and Der_compl[simp]:   "Der a (-A) = - Der a A"
+  and Der_Union[simp]:   "Der a (Union M) = Union(Der a ` M)"
+  and Der_UN[simp]:      "Der a (UN x:I. S x) = (UN x:I. Der a (S x))"
+by (auto simp: Der_def)
 
 lemma Der_star[simp]:
   shows "Der c (A\<star>) = (Der c A) ;; A\<star>"
@@ -103,6 +144,13 @@
   finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" .
 qed
 
+lemma Der_pow[simp]:
+  shows "Der c (A ^^ n) = (if n = 0 then {} else (Der c A) ;; (A ^^ (n - 1)))"
+  apply(induct n arbitrary: A)
+   apply(auto simp add: Cons_eq_append_conv)
+  by (metis Suc_pred concI_if_Nil2 conc_assoc conc_pow_comm lang_pow.simps(2))
+
+
 lemma Star_concat:
   assumes "\<forall>s \<in> set ss. s \<in> A"  
   shows "concat ss \<in> A\<star>"
@@ -119,6 +167,7 @@
 
 
 
+
 section \<open>Regular Expressions\<close>
 
 datatype rexp =
@@ -128,6 +177,7 @@
 | SEQ rexp rexp
 | ALT rexp rexp
 | STAR rexp
+| NTIMES rexp nat
 
 section \<open>Semantics of Regular Expressions\<close>
  
@@ -140,7 +190,7 @@
 | "L (SEQ r1 r2) = (L r1) ;; (L r2)"
 | "L (ALT r1 r2) = (L r1) \<union> (L r2)"
 | "L (STAR r) = (L r)\<star>"
-
+| "L (NTIMES r n) = (L r) ^^ n"
 
 section \<open>Nullable, Derivatives\<close>
 
@@ -153,7 +203,7 @@
 | "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
 | "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
 | "nullable (STAR r) = True"
-
+| "nullable (NTIMES r n) = (if n = 0 then True else nullable r)"
 
 fun
  der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
@@ -167,6 +217,8 @@
       then ALT (SEQ (der c r1) r2) (der c r2)
       else SEQ (der c r1) r2)"
 | "der c (STAR r) = SEQ (der c r) (STAR r)"
+| "der c (NTIMES r n) = (if n = 0 then ZERO else SEQ (der c r) (NTIMES r (n - 1)))"
+
 
 fun 
  ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
@@ -175,13 +227,23 @@
 | "ders (c # s) r = ders s (der c r)"
 
 
+lemma pow_empty_iff:
+  shows "[] \<in> (L r) ^^ n \<longleftrightarrow> (if n = 0 then True else [] \<in> (L r))"
+  by (induct n) (auto simp add: Sequ_def)
+
 lemma nullable_correctness:
   shows "nullable r  \<longleftrightarrow> [] \<in> (L r)"
-by (induct r) (auto simp add: Sequ_def) 
+  by (induct r) (auto simp add: Sequ_def pow_empty_iff) 
 
 lemma der_correctness:
   shows "L (der c r) = Der c (L r)"
-by (induct r) (simp_all add: nullable_correctness)
+  apply (induct r) 
+        apply(auto simp add: nullable_correctness Sequ_def)
+  using Der_def apply force
+  using Der_def apply auto[1]
+  apply (smt (verit, ccfv_SIG) Der_def append_eq_Cons_conv mem_Collect_eq)
+  using Der_def apply force
+  using Der_Sequ Sequ_def by auto
 
 lemma ders_correctness:
   shows "L (ders s r) = Ders s (L r)"
@@ -197,40 +259,4 @@
   by (simp add: ders_append)
 
 
-(*
-datatype ctxt = 
-    SeqC rexp bool
-  | AltCL rexp
-  | AltCH rexp 
-  | StarC rexp 
-
-function
-     down :: "char \<Rightarrow> rexp \<Rightarrow> ctxt list \<Rightarrow> rexp * ctxt list"
-and  up :: "char \<Rightarrow> rexp \<Rightarrow> ctxt list \<Rightarrow> rexp * ctxt list"
-where
-  "down c (SEQ r1 r2) ctxts =
-     (if (nullable r1) then down c r1 (SeqC r2 True # ctxts) 
-      else down c r1 (SeqC r2 False # ctxts))"
-| "down c (CH d) ctxts = 
-     (if c = d then up c ONE ctxts else up c ZERO ctxts)"
-| "down c ONE ctxts = up c ZERO ctxts"
-| "down c ZERO ctxts = up c ZERO ctxts"
-| "down c (ALT r1 r2) ctxts = down c r1 (AltCH r2 # ctxts)"
-| "down c (STAR r1) ctxts = down c r1 (StarC r1 # ctxts)"
-| "up c r [] = (r, [])"
-| "up c r (SeqC r2 False # ctxts) = up c (SEQ r r2) ctxts"
-| "up c r (SeqC r2 True # ctxts) = down c r2 (AltCL (SEQ r r2) # ctxts)"
-| "up c r (AltCL r1 # ctxts) = up c (ALT r1 r) ctxts"
-| "up c r (AltCH r2 # ctxts) = down c r2 (AltCL r # ctxts)"
-| "up c r (StarC r1 # ctxts) = up c (SEQ r (STAR r1)) ctxts"
-  apply(pat_completeness)
-  apply(auto)
-  done
-
-termination
-  sorry
-
-*)
-
-
 end
\ No newline at end of file