--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/BasicIdentities.thy Wed Mar 09 17:33:08 2022 +0000
@@ -0,0 +1,441 @@
+theory BasicIdentities imports
+"Lexer" "PDerivs"
+begin
+
+datatype rrexp =
+ RZERO
+| RONE
+| RCHAR char
+| RSEQ rrexp rrexp
+| RALTS "rrexp list"
+| RSTAR rrexp
+
+abbreviation
+ "RALT r1 r2 \<equiv> RALTS [r1, r2]"
+
+
+
+fun
+ rnullable :: "rrexp \<Rightarrow> bool"
+where
+ "rnullable (RZERO) = False"
+| "rnullable (RONE ) = True"
+| "rnullable (RCHAR c) = False"
+| "rnullable (RALTS rs) = (\<exists>r \<in> set rs. rnullable r)"
+| "rnullable (RSEQ r1 r2) = (rnullable r1 \<and> rnullable r2)"
+| "rnullable (RSTAR r) = True"
+
+
+fun
+ rder :: "char \<Rightarrow> rrexp \<Rightarrow> rrexp"
+where
+ "rder c (RZERO) = RZERO"
+| "rder c (RONE) = RZERO"
+| "rder c (RCHAR d) = (if c = d then RONE else RZERO)"
+| "rder c (RALTS rs) = RALTS (map (rder c) rs)"
+| "rder c (RSEQ r1 r2) =
+ (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)"
+
+
+fun
+ rders :: "rrexp \<Rightarrow> string \<Rightarrow> rrexp"
+where
+ "rders r [] = r"
+| "rders r (c#s) = rders (rder c r) s"
+
+fun rdistinct :: "'a list \<Rightarrow>'a set \<Rightarrow> 'a list"
+ where
+ "rdistinct [] acc = []"
+| "rdistinct (x#xs) acc =
+ (if x \<in> acc then rdistinct xs acc
+ else x # (rdistinct xs ({x} \<union> acc)))"
+
+
+
+
+
+fun rflts :: "rrexp list \<Rightarrow> rrexp list"
+ where
+ "rflts [] = []"
+| "rflts (RZERO # rs) = rflts rs"
+| "rflts ((RALTS rs1) # rs) = rs1 @ rflts rs"
+| "rflts (r1 # rs) = r1 # rflts rs"
+
+
+fun rsimp_ALTs :: " rrexp list \<Rightarrow> rrexp"
+ where
+ "rsimp_ALTs [] = RZERO"
+| "rsimp_ALTs [r] = r"
+| "rsimp_ALTs rs = RALTS rs"
+
+fun rsimp_SEQ :: " rrexp \<Rightarrow> rrexp \<Rightarrow> rrexp"
+ where
+ "rsimp_SEQ RZERO _ = RZERO"
+| "rsimp_SEQ _ RZERO = RZERO"
+| "rsimp_SEQ RONE r2 = r2"
+| "rsimp_SEQ r1 r2 = RSEQ r1 r2"
+
+
+fun rsimp :: "rrexp \<Rightarrow> rrexp"
+ where
+ "rsimp (RSEQ r1 r2) = rsimp_SEQ (rsimp r1) (rsimp r2)"
+| "rsimp (RALTS rs) = rsimp_ALTs (rdistinct (rflts (map rsimp rs)) {}) "
+| "rsimp r = r"
+
+
+fun
+ rders_simp :: "rrexp \<Rightarrow> string \<Rightarrow> rrexp"
+where
+ "rders_simp r [] = r"
+| "rders_simp r (c#s) = rders_simp (rsimp (rder c r)) s"
+
+fun rsize :: "rrexp \<Rightarrow> nat" where
+ "rsize RZERO = 1"
+| "rsize (RONE) = 1"
+| "rsize (RCHAR c) = 1"
+| "rsize (RALTS rs) = Suc (sum_list (map rsize rs))"
+| "rsize (RSEQ r1 r2) = Suc (rsize r1 + rsize r2)"
+| "rsize (RSTAR r) = Suc (rsize r)"
+
+
+lemma rder_rsimp_ALTs_commute:
+ shows " (rder x (rsimp_ALTs rs)) = rsimp_ALTs (map (rder x) rs)"
+ apply(induct rs)
+ apply simp
+ apply(case_tac rs)
+ apply simp
+ apply auto
+ done
+
+
+lemma rsimp_aalts_smaller:
+ shows "rsize (rsimp_ALTs rs) \<le> rsize (RALTS rs)"
+ apply(induct rs)
+ apply simp
+ apply simp
+ apply(case_tac "rs = []")
+ apply simp
+ apply(subgoal_tac "\<exists>rsp ap. rs = ap # rsp")
+ apply(erule exE)+
+ apply simp
+ apply simp
+ by(meson neq_Nil_conv)
+
+
+
+
+
+lemma rSEQ_mono:
+ shows "rsize (rsimp_SEQ r1 r2) \<le>rsize ( RSEQ r1 r2)"
+ apply auto
+ apply(induct r1)
+ apply auto
+ apply(case_tac "r2")
+ apply simp_all
+ apply(case_tac r2)
+ apply simp_all
+ apply(case_tac r2)
+ apply simp_all
+ apply(case_tac r2)
+ apply simp_all
+ apply(case_tac r2)
+ apply simp_all
+ done
+
+lemma ralts_cap_mono:
+ shows "rsize (RALTS rs) \<le> Suc ( sum_list (map rsize rs)) "
+ by simp
+
+lemma rflts_def_idiot:
+ shows "\<lbrakk> a \<noteq> RZERO; \<nexists>rs1. a = RALTS rs1\<rbrakk>
+ \<Longrightarrow> rflts (a # rs) = a # rflts rs"
+ apply(case_tac a)
+ apply simp_all
+ done
+
+
+lemma rflts_mono:
+ shows "sum_list (map rsize (rflts rs))\<le> sum_list (map rsize rs)"
+ apply(induct rs)
+ apply simp
+ apply(case_tac "a = RZERO")
+ apply simp
+ apply(case_tac "\<exists>rs1. a = RALTS rs1")
+ apply(erule exE)
+ apply simp
+ apply(subgoal_tac "rflts (a # rs) = a # (rflts rs)")
+ prefer 2
+ using rflts_def_idiot apply blast
+ apply simp
+ done
+
+lemma rdistinct_smaller: shows "sum_list (map rsize (rdistinct rs ss)) \<le>
+sum_list (map rsize rs )"
+ apply (induct rs arbitrary: ss)
+ apply simp
+ by (simp add: trans_le_add2)
+
+lemma rdistinct_phi_smaller: "sum_list (map rsize (rdistinct rs {})) \<le> sum_list (map rsize rs)"
+ by (simp add: rdistinct_smaller)
+
+
+lemma rsimp_alts_mono :
+ shows "\<And>x. (\<And>xa. xa \<in> set x \<Longrightarrow> rsize (rsimp xa) \<le> rsize xa) \<Longrightarrow>
+rsize (rsimp_ALTs (rdistinct (rflts (map rsimp x)) {})) \<le> Suc (sum_list (map rsize x))"
+ apply(subgoal_tac "rsize (rsimp_ALTs (rdistinct (rflts (map rsimp x)) {} ))
+ \<le> rsize (RALTS (rdistinct (rflts (map rsimp x)) {} ))")
+ prefer 2
+ using rsimp_aalts_smaller apply auto[1]
+ apply(subgoal_tac "rsize (RALTS (rdistinct (rflts (map rsimp x)) {})) \<le>Suc( sum_list (map rsize (rdistinct (rflts (map rsimp x)) {})))")
+ prefer 2
+ using ralts_cap_mono apply blast
+ apply(subgoal_tac "sum_list (map rsize (rdistinct (rflts (map rsimp x)) {})) \<le>
+ sum_list (map rsize ( (rflts (map rsimp x))))")
+ prefer 2
+ using rdistinct_smaller apply presburger
+ apply(subgoal_tac "sum_list (map rsize (rflts (map rsimp x))) \<le>
+ sum_list (map rsize (map rsimp x))")
+ prefer 2
+ using rflts_mono apply blast
+ apply(subgoal_tac "sum_list (map rsize (map rsimp x)) \<le> sum_list (map rsize x)")
+ prefer 2
+
+ apply (simp add: sum_list_mono)
+ by linarith
+
+
+
+
+
+lemma rsimp_mono:
+ shows "rsize (rsimp r) \<le> rsize r"
+ apply(induct r)
+ apply simp_all
+ apply(subgoal_tac "rsize (rsimp_SEQ (rsimp r1) (rsimp r2)) \<le> rsize (RSEQ (rsimp r1) (rsimp r2))")
+ apply force
+ using rSEQ_mono
+ apply presburger
+ using rsimp_alts_mono by auto
+
+lemma idiot:
+ shows "rsimp_SEQ RONE r = r"
+ apply(case_tac r)
+ apply simp_all
+ done
+
+lemma no_alt_short_list_after_simp:
+ shows "RALTS rs = rsimp r \<Longrightarrow> rsimp_ALTs rs = RALTS rs"
+ sorry
+
+lemma no_further_dB_after_simp:
+ shows "RALTS rs = rsimp r \<Longrightarrow> rdistinct rs {} = rs"
+ sorry
+
+
+lemma idiot2:
+ shows " \<lbrakk>r1 \<noteq> RZERO; r1 \<noteq> RONE;r2 \<noteq> RZERO\<rbrakk>
+ \<Longrightarrow> rsimp_SEQ r1 r2 = RSEQ r1 r2"
+ apply(case_tac r1)
+ apply(case_tac r2)
+ apply simp_all
+ apply(case_tac r2)
+ apply simp_all
+ apply(case_tac r2)
+ apply simp_all
+ apply(case_tac r2)
+ apply simp_all
+ apply(case_tac r2)
+ apply simp_all
+ done
+
+lemma rders__onechar:
+ shows " (rders_simp r [c]) = (rsimp (rders r [c]))"
+ by simp
+
+lemma rders_append:
+ "rders c (s1 @ s2) = rders (rders c s1) s2"
+ apply(induct s1 arbitrary: c s2)
+ apply(simp_all)
+ done
+
+lemma rders_simp_append:
+ "rders_simp c (s1 @ s2) = rders_simp (rders_simp c s1) s2"
+ apply(induct s1 arbitrary: c s2)
+ apply(simp_all)
+ done
+
+lemma inside_simp_removal:
+ shows " rsimp (rder x (rsimp r)) = rsimp (rder x r)"
+ sorry
+
+lemma set_related_list:
+ shows "distinct rs \<Longrightarrow> length rs = card (set rs)"
+ by (simp add: distinct_card)
+(*this section deals with the property of distinctBy: creates a list without duplicates*)
+lemma rdistinct_never_added_twice:
+ shows "rdistinct (a # rs) {a} = rdistinct rs {a}"
+ by force
+
+
+lemma rdistinct_does_the_job:
+ shows "distinct (rdistinct rs s)"
+ apply(induct rs arbitrary: s)
+ apply simp
+ apply simp
+ sorry
+
+lemma rders_simp_same_simpders:
+ shows "s \<noteq> [] \<Longrightarrow> rders_simp r s = rsimp (rders r s)"
+ apply(induct s rule: rev_induct)
+ apply simp
+ apply(case_tac "xs = []")
+ apply simp
+ apply(simp add: rders_append rders_simp_append)
+ using inside_simp_removal by blast
+
+lemma simp_helps_der_pierce:
+ shows " rsimp
+ (rder x
+ (rsimp_ALTs rs)) =
+ rsimp
+ (rsimp_ALTs
+ (map (rder x )
+ rs
+ )
+ )"
+ sorry
+
+
+lemma rders_simp_one_char:
+ shows "rders_simp r [c] = rsimp (rder c r)"
+ apply auto
+ done
+
+lemma rsimp_idem:
+ shows "rsimp (rsimp r) = rsimp r"
+ sorry
+
+corollary rsimp_inner_idem1:
+ shows "rsimp r = RSEQ r1 r2 \<Longrightarrow> rsimp r1 = r1 \<and> rsimp r2 = r2"
+
+ sorry
+
+corollary rsimp_inner_idem2:
+ shows "rsimp r = RALTS rs \<Longrightarrow> \<forall>r' \<in> (set rs). rsimp r' = r'"
+ sorry
+
+corollary rsimp_inner_idem3:
+ shows "rsimp r = RALTS rs \<Longrightarrow> map rsimp rs = rs"
+ by (meson map_idI rsimp_inner_idem2)
+
+corollary rsimp_inner_idem4:
+ shows "rsimp r = RALTS rs \<Longrightarrow> flts rs = rs"
+ sorry
+
+
+lemma head_one_more_simp:
+ shows "map rsimp (r # rs) = map rsimp (( rsimp r) # rs)"
+ by (simp add: rsimp_idem)
+
+lemma head_one_more_dersimp:
+ shows "map rsimp ((rder x (rders_simp r s) # rs)) = map rsimp ((rders_simp r (s@[x]) ) # rs)"
+ using head_one_more_simp rders_simp_append rders_simp_one_char by presburger
+
+
+
+
+lemma ders_simp_nullability:
+ shows "rnullable (rders r s) = rnullable (rders_simp r s)"
+ sorry
+
+lemma first_elem_seqder:
+ shows "\<not>rnullable r1p \<Longrightarrow> map rsimp (rder x (RSEQ r1p r2)
+ # rs) = map rsimp ((RSEQ (rder x r1p) r2) # rs) "
+ by auto
+
+lemma first_elem_seqder1:
+ shows "\<not>rnullable (rders_simp r xs) \<Longrightarrow> map rsimp ( (rder x (RSEQ (rders_simp r xs) r2)) # rs) =
+ map rsimp ( (RSEQ (rsimp (rder x (rders_simp r xs))) r2) # rs)"
+ by (simp add: rsimp_idem)
+
+lemma first_elem_seqdersimps:
+ shows "\<not>rnullable (rders_simp r xs) \<Longrightarrow> map rsimp ( (rder x (RSEQ (rders_simp r xs) r2)) # rs) =
+ map rsimp ( (RSEQ (rders_simp r (xs @ [x])) r2) # rs)"
+ using first_elem_seqder1 rders_simp_append by auto
+
+
+
+
+
+
+lemma seq_ders_closed_form1:
+ shows "\<exists>Ss. rders_simp (RSEQ r1 r2) [c] = rsimp (RALTS ((RSEQ (rders_simp r1 [c]) r2) #
+(map ( rders_simp r2 ) Ss)))"
+ apply(case_tac "rnullable r1")
+ apply(subgoal_tac " rders_simp (RSEQ r1 r2) [c] =
+rsimp (RALTS ((RSEQ (rders_simp r1 [c]) r2) # (map (rders_simp r2) [[c]])))")
+ prefer 2
+ apply (simp add: rsimp_idem)
+ apply(rule_tac x = "[[c]]" in exI)
+ apply simp
+ apply(subgoal_tac " rders_simp (RSEQ r1 r2) [c] =
+rsimp (RALTS ((RSEQ (rders_simp r1 [c]) r2) # (map (rders_simp r2) [])))")
+ apply blast
+ apply(simp add: rsimp_idem)
+ sorry
+
+
+
+
+
+
+
+
+lemma simp_flatten2:
+ shows "rsimp (RALTS (r # [RALTS rs])) = rsimp (RALTS (r # rs))"
+ sorry
+
+
+lemma simp_flatten:
+ shows "rsimp (RALTS ((RALTS rsa) # rsb)) = rsimp (RALTS (rsa @ rsb))"
+
+ sorry
+
+
+
+fun vsuf :: "char list \<Rightarrow> rrexp \<Rightarrow> char list list" where
+"vsuf [] _ = []"
+|"vsuf (c#cs) r1 = (if (rnullable r1) then (vsuf cs (rder c r1)) @ [c # cs]
+ else (vsuf cs (rder c r1))
+ ) "
+
+
+
+
+
+
+fun star_update :: "char \<Rightarrow> rrexp \<Rightarrow> char list list \<Rightarrow> char list list" where
+"star_update c r [] = []"
+|"star_update c r (s # Ss) = (if (rnullable (rders_simp r s))
+ then (s@[c]) # [c] # (star_update c r Ss)
+ else (s@[c]) # (star_update c r Ss) )"
+
+fun star_updates :: "char list \<Rightarrow> rrexp \<Rightarrow> char list list \<Rightarrow> char list list"
+ where
+"star_updates [] r Ss = Ss"
+| "star_updates (c # cs) r Ss = star_updates cs r (star_update c r Ss)"
+
+
+
+
+
+
+(*some basic facts about rsimp*)
+
+
+
+
+end
\ No newline at end of file