diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/FBound.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/FBound.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,180 @@ + +theory FBound + imports "BlexerSimp" "ClosedFormsBounds" +begin + +fun distinctBy :: "'a list \ ('a \ 'b) \ 'b set \ 'a list" + where + "distinctBy [] f acc = []" +| "distinctBy (x#xs) f acc = + (if (f x) \ acc then distinctBy xs f acc + else x # (distinctBy xs f ({f x} \ acc)))" + +fun rerase :: "arexp \ rrexp" +where + "rerase AZERO = RZERO" +| "rerase (AONE _) = RONE" +| "rerase (ACHAR _ c) = RCHAR c" +| "rerase (AALTs bs rs) = RALTS (map rerase rs)" +| "rerase (ASEQ _ r1 r2) = RSEQ (rerase r1) (rerase r2)" +| "rerase (ASTAR _ r) = RSTAR (rerase r)" + +lemma eq1_rerase: + shows "x ~1 y \ (rerase x) = (rerase y)" + apply(induct x y rule: eq1.induct) + apply(auto) + done + + +lemma distinctBy_distinctWith: + shows "distinctBy xs f (f ` acc) = distinctWith xs (\x y. f x = f y) acc" + apply(induct xs arbitrary: acc) + apply(auto) + by (metis image_insert) + +lemma distinctBy_distinctWith2: + shows "distinctBy xs rerase {} = distinctWith xs eq1 {}" + apply(subst distinctBy_distinctWith[of _ _ "{}", simplified]) + using eq1_rerase by presburger + +lemma asize_rsize: + shows "rsize (rerase r) = asize r" + apply(induct r rule: rerase.induct) + apply(auto) + apply (metis (mono_tags, lifting) comp_apply map_eq_conv) + done + +lemma rerase_fuse: + shows "rerase (fuse bs r) = rerase r" + apply(induct r) + apply simp+ + done + +lemma rerase_bsimp_ASEQ: + shows "rerase (bsimp_ASEQ x1 a1 a2) = rsimp_SEQ (rerase a1) (rerase a2)" + apply(induct x1 a1 a2 rule: bsimp_ASEQ.induct) + apply(auto) + done + +lemma rerase_bsimp_AALTs: + shows "rerase (bsimp_AALTs bs rs) = rsimp_ALTs (map rerase rs)" + apply(induct bs rs rule: bsimp_AALTs.induct) + apply(auto simp add: rerase_fuse) + done + +fun anonalt :: "arexp \ bool" + where + "anonalt (AALTs bs2 rs) = False" +| "anonalt r = True" + + +fun agood :: "arexp \ bool" where + "agood AZERO = False" +| "agood (AONE cs) = True" +| "agood (ACHAR cs c) = True" +| "agood (AALTs cs []) = False" +| "agood (AALTs cs [r]) = False" +| "agood (AALTs cs (r1#r2#rs)) = (distinct (map rerase (r1 # r2 # rs)) \(\r' \ set (r1#r2#rs). agood r' \ anonalt r'))" +| "agood (ASEQ _ AZERO _) = False" +| "agood (ASEQ _ (AONE _) _) = False" +| "agood (ASEQ _ _ AZERO) = False" +| "agood (ASEQ cs r1 r2) = (agood r1 \ agood r2)" +| "agood (ASTAR cs r) = True" + + +fun anonnested :: "arexp \ bool" + where + "anonnested (AALTs bs2 []) = True" +| "anonnested (AALTs bs2 ((AALTs bs1 rs1) # rs2)) = False" +| "anonnested (AALTs bs2 (r # rs2)) = anonnested (AALTs bs2 rs2)" +| "anonnested r = True" + + +lemma asize0: + shows "0 < asize r" + apply(induct r) + apply(auto) + done + +lemma rnullable: + shows "rnullable (rerase r) = bnullable r" + apply(induct r rule: rerase.induct) + apply(auto) + done + +lemma rder_bder_rerase: + shows "rder c (rerase r ) = rerase (bder c r)" + apply (induct r) + apply (auto) + using rerase_fuse apply presburger + using rnullable apply blast + using rnullable by blast + +lemma rerase_map_bsimp: + assumes "\ r. r \ set rs \ rerase (bsimp r) = (rsimp \ rerase) r" + shows "map rerase (map bsimp rs) = map (rsimp \ rerase) rs" + using assms + apply(induct rs) + by simp_all + + +lemma rerase_flts: + shows "map rerase (flts rs) = rflts (map rerase rs)" + apply(induct rs rule: flts.induct) + apply(auto simp add: rerase_fuse) + done + +lemma rerase_dB: + shows "map rerase (distinctBy rs rerase acc) = rdistinct (map rerase rs) acc" + apply(induct rs arbitrary: acc) + apply simp+ + done + +lemma rerase_earlier_later_same: + assumes " \r. r \ set rs \ rerase (bsimp r) = rsimp (rerase r)" + shows " (map rerase (distinctBy (flts (map bsimp rs)) rerase {})) = + (rdistinct (rflts (map (rsimp \ rerase) rs)) {})" + apply(subst rerase_dB) + apply(subst rerase_flts) + apply(subst rerase_map_bsimp) + apply auto + using assms + apply simp + done + +lemma bsimp_rerase: + shows "rerase (bsimp a) = rsimp (rerase a)" + apply(induct a rule: bsimp.induct) + apply(auto) + using rerase_bsimp_ASEQ apply presburger + using distinctBy_distinctWith2 rerase_bsimp_AALTs rerase_earlier_later_same by fastforce + +lemma rders_simp_size: + shows "rders_simp (rerase r) s = rerase (bders_simp r s)" + apply(induct s rule: rev_induct) + apply simp + by (simp add: bders_simp_append rder_bder_rerase rders_simp_append bsimp_rerase) + + +corollary aders_simp_finiteness: + assumes "\N. \s. rsize (rders_simp (rerase r) s) \ N" + shows " \N. \s. asize (bders_simp r s) \ N" +proof - + from assms obtain N where "\s. rsize (rders_simp (rerase r) s) \ N" + by blast + then have "\s. rsize (rerase (bders_simp r s)) \ N" + by (simp add: rders_simp_size) + then have "\s. asize (bders_simp r s) \ N" + by (simp add: asize_rsize) + then show "\N. \s. asize (bders_simp r s) \ N" by blast +qed + +theorem annotated_size_bound: + shows "\N. \s. asize (bders_simp r s) \ N" + apply(insert aders_simp_finiteness) + by (simp add: rders_simp_bounded) + + +unused_thms + +end