diff --git a/thys/Earley_Parser/CFG.thy b/thys/Earley_Parser/CFG.thy new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/CFG.thy @@ -0,0 +1,67 @@ +theory CFG + imports Main +begin + +section \Adjusted content from AFP/LocalLexing\ + +type_synonym 'a rule = "'a \ 'a list" + +type_synonym 'a rules = "'a rule list" + +type_synonym 'a sentence = "'a list" + +datatype 'a cfg = + CFG (\ : "'a list") (\ : "'a list") (\ : "'a rules") (\ : "'a") + +definition disjunct_symbols :: "'a cfg \ bool" where + "disjunct_symbols \ \ set (\ \) \ set (\ \) = {}" + +definition valid_startsymbol :: "'a cfg \ bool" where + "valid_startsymbol \ \ \ \ \ set (\ \)" + +definition valid_rules :: "'a cfg \ bool" where + "valid_rules \ \ \(N, \) \ set (\ \). N \ set (\ \) \ (\s \ set \. s \ set (\ \) \ set (\ \))" + +definition distinct_rules :: "'a cfg \ bool" where + "distinct_rules \ \ distinct (\ \)" + +definition wf_\ :: "'a cfg \ bool" where + "wf_\ \ \ disjunct_symbols \ \ valid_startsymbol \ \ valid_rules \ \ distinct_rules \" + +lemmas wf_\_defs = wf_\_def valid_rules_def valid_startsymbol_def disjunct_symbols_def distinct_rules_def + +definition is_terminal :: "'a cfg \ 'a \ bool" where + "is_terminal \ x \ x \ set (\ \)" + +definition is_nonterminal :: "'a cfg \ 'a \ bool" where + "is_nonterminal \ x \ x \ set (\ \)" + +definition is_symbol :: "'a cfg \ 'a \ bool" where + "is_symbol \ x \ is_terminal \ x \ is_nonterminal \ x" + +definition wf_sentence :: "'a cfg \ 'a sentence \ bool" where + "wf_sentence \ \ \ \x \ set \. is_symbol \ x" + +lemma is_nonterminal_startsymbol: + "wf_\ \ \ is_nonterminal \ (\ \)" + by (simp add: is_nonterminal_def wf_\_defs) + +definition is_word :: "'a cfg \ 'a sentence \ bool" where + "is_word \ \ \ \x \ set \. is_terminal \ x" + +definition derives1 :: "'a cfg \ 'a sentence \ 'a sentence \ bool" where + "derives1 \ u v \ \ x y N \. + u = x @ [N] @ y \ + v = x @ \ @ y \ + (N, \) \ set (\ \)" + +definition derivations1 :: "'a cfg \ ('a sentence \ 'a sentence) set" where + "derivations1 \ \ { (u,v) | u v. derives1 \ u v }" + +definition derivations :: "'a cfg \ ('a sentence \ 'a sentence) set" where + "derivations \ \ (derivations1 \)^*" + +definition derives :: "'a cfg \ 'a sentence \ 'a sentence \ bool" where + "derives \ u v \ ((u, v) \ derivations \)" + +end \ No newline at end of file diff --git a/thys/Earley_Parser/Derivations.thy b/thys/Earley_Parser/Derivations.thy new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/Derivations.thy @@ -0,0 +1,255 @@ +theory Derivations + imports + CFG +begin + +section \Adjusted content from AFP/LocalLexing\ + +type_synonym 'a derivation = "(nat \ 'a rule) list" + +lemma is_word_empty: "is_word \ []" by (auto simp add: is_word_def) + +lemma derives1_implies_derives[simp]: + "derives1 \ a b \ derives \ a b" + by (auto simp add: derives_def derivations_def derivations1_def) + +lemma derives_trans: + "derives \ a b \ derives \ b c \ derives \ a c" + by (auto simp add: derives_def derivations_def) + +lemma derives1_eq_derivations1: + "derives1 \ x y = ((x, y) \ derivations1 \)" + by (simp add: derivations1_def) + +lemma derives_induct[consumes 1, case_names Base Step]: + assumes derives: "derives \ a b" + assumes Pa: "P a" + assumes induct: "\y z. derives \ a y \ derives1 \ y z \ P y \ P z" + shows "P b" +proof - + note rtrancl_lemma = rtrancl_induct[where a = a and b = b and r = "derivations1 \" and P=P] + from derives Pa induct rtrancl_lemma show "P b" + by (metis derives_def derivations_def derives1_eq_derivations1) +qed + +definition Derives1 :: "'a cfg \ 'a sentence \ nat \ 'a rule \ 'a sentence \ bool" where + "Derives1 \ u i r v \ \ x y N \. + u = x @ [N] @ y \ + v = x @ \ @ y \ + (N, \) \ set (\ \) \ r = (N, \) \ i = length x" + +lemma Derives1_split: + "Derives1 \ u i r v \ \ x y. u = x @ [fst r] @ y \ v = x @ (snd r) @ y \ length x = i" + by (metis Derives1_def fst_conv snd_conv) + +lemma Derives1_implies_derives1: "Derives1 \ u i r v \ derives1 \ u v" + by (auto simp add: Derives1_def derives1_def) + +lemma derives1_implies_Derives1: "derives1 \ u v \ \ i r. Derives1 \ u i r v" + by (auto simp add: Derives1_def derives1_def) + +fun Derivation :: "'a cfg \ 'a sentence \ 'a derivation \ 'a sentence \ bool" where + "Derivation _ a [] b = (a = b)" +| "Derivation \ a (d#D) b = (\ x. Derives1 \ a (fst d) (snd d) x \ Derivation \ x D b)" + +lemma Derivation_implies_derives: "Derivation \ a D b \ derives \ a b" +proof (induct D arbitrary: a b) + case Nil thus ?case + by (simp add: derives_def derivations_def) +next + case (Cons d D) + note ihyps = this + from ihyps have "\ x. Derives1 \ a (fst d) (snd d) x \ Derivation \ x D b" by auto + then obtain x where "Derives1 \ a (fst d) (snd d) x" and xb: "Derivation \ x D b" by blast + with Derives1_implies_derives1 have d1: "derives \ a x" by fastforce + from ihyps xb have d2:"derives \ x b" by simp + show "derives \ a b" by (rule derives_trans[OF d1 d2]) +qed + +lemma Derivation_Derives1: "Derivation \ a S y \ Derives1 \ y i r z \ Derivation \ a (S@[(i,r)]) z" +proof (induct S arbitrary: a y z i r) + case Nil thus ?case by simp +next + case (Cons s S) thus ?case + by (metis Derivation.simps(2) append_Cons) +qed + +lemma derives_implies_Derivation: "derives \ a b \ \ D. Derivation \ a D b" +proof (induct rule: derives_induct) + case Base + show ?case by (rule exI[where x="[]"], simp) +next + case (Step y z) + note ihyps = this + from ihyps obtain D where ay: "Derivation \ a D y" by blast + from ihyps derives1_implies_Derives1 obtain i r where yz: "Derives1 \ y i r z" by blast + from Derivation_Derives1[OF ay yz] show ?case by auto +qed + +lemma rule_nonterminal_type[simp]: "wf_\ \ \ (N, \) \ set (\ \) \ is_nonterminal \ N" + by (auto simp add: is_nonterminal_def wf_\_defs) + +lemma Derives1_rule [elim]: "Derives1 \ a i r b \ r \ set (\ \)" + using Derives1_def by metis + +lemma is_terminal_nonterminal: "wf_\ \ \ is_terminal \ x \ is_nonterminal \ x \ False" + by (auto simp: wf_\_defs disjoint_iff is_nonterminal_def is_terminal_def) + +lemma is_word_is_terminal: "i < length u \ is_word \ u \ is_terminal \ (u ! i)" + using is_word_def by force + +lemma Derivation_append: "Derivation \ a (D@E) c = (\ b. Derivation \ a D b \ Derivation \ b E c)" + by (induct D arbitrary: a c E) auto + +lemma Derivation_implies_append: + "Derivation \ a D b \ Derivation \ b E c \ Derivation \ a (D@E) c" + using Derivation_append by blast + + +section \Additional derivation lemmas\ + +lemma Derives1_prepend: + assumes "Derives1 \ u i r v" + shows "Derives1 \ (w@u) (i + length w) r (w@v)" +proof - + obtain x y N \ where *: + "u = x @ [N] @ y" "v = x @ \ @ y" + "(N, \) \ set (\ \)" "r = (N, \)" "i = length x" + using assms Derives1_def by (smt (verit)) + hence "w@u = w @ x @ [N] @ y" "w@v = w @ x @ \ @ y" + by auto + thus ?thesis + unfolding Derives1_def using * + apply (rule_tac exI[where x="w@x"]) + apply (rule_tac exI[where x="y"]) + by simp +qed + +lemma Derivation_prepend: + "Derivation \ b D b' \ Derivation \ (a@b) (map (\(i, r). (i + length a, r)) D) (a@b')" + using Derives1_prepend by (induction D arbitrary: b b') (auto, fast) + +lemma Derives1_append: + assumes "Derives1 \ u i r v" + shows "Derives1 \ (u@w) i r (v@w)" +proof - + obtain x y N \ where *: + "u = x @ [N] @ y" "v = x @ \ @ y" + "(N, \) \ set (\ \)" "r = (N, \)" "i = length x" + using assms Derives1_def by (smt (verit)) + hence "u@w = x @ [N] @ y @ w" "v@w = x @ \ @ y @ w" + by auto + thus ?thesis + unfolding Derives1_def using * + apply (rule_tac exI[where x="x"]) + apply (rule_tac exI[where x="y@w"]) + by blast +qed + +lemma Derivation_append': + "Derivation \ a D a' \ Derivation \ (a@b) D (a'@b)" + using Derives1_append by (induction D arbitrary: a a') (auto, fast) + +lemma Derivation_append_rewrite: + assumes "Derivation \ a D (b @ c @ d) " "Derivation \ c E c'" + shows "\F. Derivation \ a F (b @ c' @ d)" + using assms Derivation_append' Derivation_prepend Derivation_implies_append by fast + +lemma derives1_if_valid_rule: + "(N, \) \ set (\ \) \ derives1 \ [N] \" + unfolding derives1_def + apply (rule_tac exI[where x="[]"]) + apply (rule_tac exI[where x="[]"]) + by simp + +lemma derives_if_valid_rule: + "(N, \) \ set (\ \) \ derives \ [N] \" + using derives1_if_valid_rule by fastforce + +lemma Derivation_from_empty: + "Derivation \ [] D a \ a = []" + by (cases D) (auto simp: Derives1_def) + +lemma Derivation_concat_split: + "Derivation \ (a@b) D c \ \E F a' b'. Derivation \ a E a' \ Derivation \ b F b' \ + c = a' @ b' \ length E \ length D \ length F \ length D" +proof (induction D arbitrary: a b) + case Nil + thus ?case + by (metis Derivation.simps(1) order_refl) +next + case (Cons d D) + then obtain ab where *: "Derives1 \ (a@b) (fst d) (snd d) ab" "Derivation \ ab D c" + by auto + then obtain x y N \ where #: + "a@b = x @ [N] @ y" "ab = x @ \ @ y" "(N,\) \ set (\ \)" "snd d = (N,\)" "fst d = length x" + using * unfolding Derives1_def by blast + show ?case + proof (cases "length a \ length x") + case True + hence ab_def: + "a = take (length a) x" + "b = drop (length a) x @ [N] @ y" + "ab = take (length a) x @ drop (length a) x @ \ @ y" + using #(1,2) True by (metis append_eq_append_conv_if)+ + then obtain E F a' b' where IH: + "Derivation \ (take (length a) x) E a'" + "Derivation \ (drop (length a) x @ \ @ y) F b'" + "c = a' @ b'" + "length E \ length D" + "length F \ length D" + using Cons *(2) by blast + have "Derives1 \ b (fst d - length a) (snd d) (drop (length a) x @ \ @ y)" + unfolding Derives1_def using *(1) #(3-5) ab_def(2) by (metis length_drop) + hence "Derivation \ b ((fst d - length a, snd d) # F) b'" + using IH(2) by force + moreover have "Derivation \ a E a'" + using IH(1) ab_def(1) by fastforce + ultimately show ?thesis + using IH(3-5) by fastforce + next + case False + hence a_def: "a = x @ [N] @ take (length a - length x - 1) y" + using #(1) append_eq_conv_conj[of a b "x @ [N] @ y"] take_all_iff take_append + by (metis append_Cons append_Nil diff_is_0_eq le_cases take_Cons') + hence b_def: "b = drop (length a - length x - 1) y" + using #(1) by (metis List.append.assoc append_take_drop_id same_append_eq) + have "ab = x @ \ @ take (length a - length x - 1) y @ drop (length a - length x - 1) y" + using #(2) by force + then obtain E F a' b' where IH: + "Derivation \ (x @ \ @ take (length a - length x - 1) y) E a'" + "Derivation \ (drop (length a - length x - 1) y) F b'" + "c = a' @ b'" + "length E \ length D" + "length F \ length D" + using Cons.IH[of "x @ \ @ take (length a - length x - 1) y" "drop (length a - length x - 1) y"] *(2) by auto + have "Derives1 \ a (fst d) (snd d) (x @ \ @ take (length a - length x - 1) y)" + unfolding Derives1_def using #(3-5) a_def by blast + hence "Derivation \ a ((fst d, snd d) # E) a'" + using IH(1) by fastforce + moreover have "Derivation \ b F b'" + using b_def IH(2) by blast + ultimately show ?thesis + using IH(3-5) by fastforce + qed +qed + +lemma Derivation_\1: + assumes "Derivation \ [\ \] D \" "is_word \ \" "wf_\ \" + shows "\\ E. Derivation \ \ E \ \ (\ \,\) \ set (\ \)" +proof (cases D) + case Nil + thus ?thesis + using assms is_nonterminal_startsymbol is_terminal_nonterminal by (metis Derivation.simps(1) is_word_def list.set_intros(1)) +next + case (Cons d D) + then obtain \ where "Derives1 \ [\ \] (fst d) (snd d) \" "Derivation \ \ D \" + using assms by auto + hence "(\ \, \) \ set (\ \)" + unfolding Derives1_def + by (metis List.append.right_neutral List.list.discI append_eq_Cons_conv append_is_Nil_conv nth_Cons_0 self_append_conv2) + thus ?thesis + using \Derivation \ \ D \\ by auto +qed + +end \ No newline at end of file diff --git a/thys/Earley_Parser/Earley.thy b/thys/Earley_Parser/Earley.thy new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/Earley.thy @@ -0,0 +1,493 @@ +theory Earley + imports + Derivations +begin + +section \Slices\ + +fun slice :: "nat \ nat \ 'a list \ 'a list" where + "slice _ _ [] = []" +| "slice _ 0 (x#xs) = []" +| "slice 0 (Suc b) (x#xs) = x # slice 0 b xs" +| "slice (Suc a) (Suc b) (x#xs) = slice a b xs" + +lemma slice_drop_take: + "slice a b xs = drop a (take b xs)" + by (induction a b xs rule: slice.induct) auto + +lemma slice_append_aux: + "Suc b \ c \ slice (Suc b) c (x # xs) = slice b (c-1) xs" + using Suc_le_D by fastforce + +lemma slice_concat: + "a \ b \ b \ c \ slice a b xs @ slice b c xs = slice a c xs" +proof (induction a b xs arbitrary: c rule: slice.induct) + case (3 b x xs) + then show ?case + using Suc_le_D by(fastforce simp: slice_append_aux) +qed (auto simp: slice_append_aux) + +lemma slice_concat_Ex: + "a \ c \ slice a c xs = ys @ zs \ \b. ys = slice a b xs \ zs = slice b c xs \ a \ b \ b \ c" +proof (induction a c xs arbitrary: ys zs rule: slice.induct) + case (3 b x xs) + show ?case + proof (cases ys) + case Nil + then obtain zs' where "x # slice 0 b xs = x # zs'" "x # zs' = zs" + using "3.prems"(2) by auto + thus ?thesis + using Nil by force + next + case (Cons y ys') + then obtain ys' where "x # slice 0 b xs = x # ys' @ zs" "x # ys' = ys" + using "3.prems"(2) by auto + thus ?thesis + using "3.IH"[of ys' zs] by force + qed +next + case (4 a b x xs) + thus ?case + by (auto, metis slice.simps(4) Suc_le_mono) +qed auto + +lemma slice_nth: + "a < length xs \ slice a (a+1) xs = [xs!a]" + unfolding slice_drop_take + by (metis Cons_nth_drop_Suc One_nat_def diff_add_inverse drop_take take_Suc_Cons take_eq_Nil) + +lemma slice_append_nth: + "a \ b \ b < length xs \ slice a b xs @ [xs!b] = slice a (b+1) xs" + by (metis le_add1 slice_concat slice_nth) + +lemma slice_empty: + "b \ a \ slice a b xs = []" + by (simp add: slice_drop_take) + +lemma slice_id[simp]: + "slice 0 (length xs) xs = xs" + by (simp add: slice_drop_take) + +lemma slice_singleton: + "b \ length xs \ [x] = slice a b xs \ b = a + 1" + by (induction a b xs rule: slice.induct) (auto simp: slice_drop_take) + + +section \Earley recognizer\ + +subsection \Earley items\ + +definition rule_head :: "'a rule \ 'a" where + "rule_head \ fst" + +definition rule_body :: "'a rule \ 'a list" where + "rule_body \ snd" + +datatype 'a item = + Item (item_rule: "'a rule") (item_dot : nat) (item_origin : nat) (item_end : nat) + +definition item_rule_head :: "'a item \ 'a" where + "item_rule_head x \ rule_head (item_rule x)" + +definition item_rule_body :: "'a item \ 'a sentence" where + "item_rule_body x \ rule_body (item_rule x)" + +definition item_\ :: "'a item \ 'a sentence" where + "item_\ x \ take (item_dot x) (item_rule_body x)" + +definition item_\ :: "'a item \ 'a sentence" where + "item_\ x \ drop (item_dot x) (item_rule_body x)" + +definition is_complete :: "'a item \ bool" where + "is_complete x \ item_dot x \ length (item_rule_body x)" + +definition next_symbol :: "'a item \ 'a option" where + "next_symbol x \ if is_complete x then None else Some (item_rule_body x ! item_dot x)" + +lemmas item_defs = item_rule_head_def item_rule_body_def item_\_def item_\_def rule_head_def rule_body_def + +definition is_finished :: "'a cfg \ 'a sentence \ 'a item \ bool" where + "is_finished \ \ x \ + item_rule_head x = \ \ \ + item_origin x = 0 \ + item_end x = length \ \ + is_complete x" + +definition recognizing :: "'a item set \ 'a cfg \ 'a sentence \ bool" where + "recognizing I \ \ \ \x \ I. is_finished \ \ x" + +inductive_set Earley :: "'a cfg \ 'a sentence \ 'a item set" + for \ :: "'a cfg" and \ :: "'a sentence" where + Init: "r \ set (\ \) \ fst r = \ \ \ + Item r 0 0 0 \ Earley \ \" + | Scan: "x = Item r b i j \ x \ Earley \ \ \ + \!j = a \ j < length \ \ next_symbol x = Some a \ + Item r (b + 1) i (j + 1) \ Earley \ \" + | Predict: "x = Item r b i j \ x \ Earley \ \ \ + r' \ set (\ \) \ next_symbol x = Some (rule_head r') \ + Item r' 0 j j \ Earley \ \" + | Complete: "x = Item r\<^sub>x b\<^sub>x i j \ x \ Earley \ \ \ y = Item r\<^sub>y b\<^sub>y j k \ y \ Earley \ \ \ + is_complete y \ next_symbol x = Some (item_rule_head y) \ + Item r\<^sub>x (b\<^sub>x + 1) i k \ Earley \ \" + + +subsection \Well-formedness\ + +definition wf_item :: "'a cfg \ 'a sentence => 'a item \ bool" where + "wf_item \ \ x \ + item_rule x \ set (\ \) \ + item_dot x \ length (item_rule_body x) \ + item_origin x \ item_end x \ + item_end x \ length \" + +lemma wf_Init: + assumes "r \ set (\ \)" "fst r = \ \" + shows "wf_item \ \ (Item r 0 0 0)" + using assms unfolding wf_item_def by simp + +lemma wf_Scan: + assumes "x = Item r b i j" "wf_item \ \ x" "\!j = a" "j < length \" "next_symbol x = Some a" + shows "wf_item \ \ (Item r (b + 1) i (j+1))" + using assms unfolding wf_item_def by (auto simp: item_defs is_complete_def next_symbol_def split: if_splits) + +lemma wf_Predict: + assumes "x = Item r b i j" "wf_item \ \ x" "r' \ set (\ \)" "next_symbol x = Some (rule_head r')" + shows "wf_item \ \ (Item r' 0 j j)" + using assms unfolding wf_item_def by simp + +lemma wf_Complete: + assumes "x = Item r\<^sub>x b\<^sub>x i j" "wf_item \ \ x" "y = Item r\<^sub>y b\<^sub>y j k" "wf_item \ \ y" + assumes "is_complete y" "next_symbol x = Some (item_rule_head y)" + shows "wf_item \ \ (Item r\<^sub>x (b\<^sub>x + 1) i k)" + using assms unfolding wf_item_def is_complete_def next_symbol_def item_rule_body_def + by (auto split: if_splits) + +lemma wf_Earley: + assumes "x \ Earley \ \" + shows "wf_item \ \ x" + using assms wf_Init wf_Scan wf_Predict wf_Complete + by (induction rule: Earley.induct) fast+ + + +subsection \Soundness\ + +definition sound_item :: "'a cfg \ 'a sentence \ 'a item \ bool" where + "sound_item \ \ x \ derives \ [item_rule_head x] (slice (item_origin x) (item_end x) \ @ item_\ x)" + +lemma sound_Init: + assumes "r \ set (\ \)" "fst r = \ \" + shows "sound_item \ \ (Item r 0 0 0)" +proof - + let ?x = "Item r 0 0 0" + have "(item_rule_head ?x, item_\ ?x) \ set (\ \)" + using assms(1) by (simp add: item_defs) + hence "derives \ [item_rule_head ?x] (item_\ ?x)" + using derives_if_valid_rule by metis + thus "sound_item \ \ ?x" + unfolding sound_item_def by (simp add: slice_empty) +qed + +lemma sound_Scan: + assumes "x = Item r b i j" "wf_item \ \ x" "sound_item \ \ x" + assumes "\!j = a" "j < length \" "next_symbol x = Some a" + shows "sound_item \ \ (Item r (b+1) i (j+1))" +proof - + define x' where [simp]: "x' = Item r (b+1) i (j+1)" + obtain item_\' where *: "item_\ x = a # item_\'" "item_\ x' = item_\'" + using assms(1,6) apply (auto simp: item_defs next_symbol_def is_complete_def split: if_splits) + by (metis Cons_nth_drop_Suc leI) + have "slice i j \ @ item_\ x = slice i (j+1) \ @ item_\'" + using * assms(1,2,4,5) by (auto simp: slice_append_nth wf_item_def) + moreover have "derives \ [item_rule_head x] (slice i j \ @ item_\ x)" + using assms(1,3) sound_item_def by force + ultimately show ?thesis + using assms(1) * by (auto simp: item_defs sound_item_def) +qed + +lemma sound_Predict: + assumes "x = Item r b i j" "wf_item \ \ x" "sound_item \ \ x" + assumes "r' \ set (\ \)" "next_symbol x = Some (rule_head r')" + shows "sound_item \ \ (Item r' 0 j j)" + using assms by (auto simp: sound_item_def derives_if_valid_rule slice_empty item_defs) + +lemma sound_Complete: + assumes "x = Item r\<^sub>x b\<^sub>x i j" "wf_item \ \ x" "sound_item \ \ x" + assumes "y = Item r\<^sub>y b\<^sub>y j k" "wf_item \ \ y" "sound_item \ \ y" + assumes "is_complete y" "next_symbol x = Some (item_rule_head y)" + shows "sound_item \ \ (Item r\<^sub>x (b\<^sub>x + 1) i k)" +proof - + have "derives \ [item_rule_head y] (slice j k \)" + using assms(4,6,7) by (auto simp: sound_item_def is_complete_def item_defs) + then obtain E where E: "Derivation \ [item_rule_head y] E (slice j k \)" + using derives_implies_Derivation by blast + have "derives \ [item_rule_head x] (slice i j \ @ item_\ x)" + using assms(1,3,4) by (auto simp: sound_item_def) + moreover have 0: "item_\ x = (item_rule_head y) # tl (item_\ x)" + using assms(8) apply (auto simp: next_symbol_def is_complete_def item_defs split: if_splits) + by (metis drop_eq_Nil hd_drop_conv_nth leI list.collapse) + ultimately obtain D where D: + "Derivation \ [item_rule_head x] D (slice i j \ @ [item_rule_head y] @ (tl (item_\ x)))" + using derives_implies_Derivation by (metis append_Cons append_Nil) + obtain F where F: + "Derivation \ [item_rule_head x] F (slice i j \ @ slice j k \ @ tl (item_\ x))" + using Derivation_append_rewrite D E by blast + moreover have "i \ j" + using assms(1,2) wf_item_def by force + moreover have "j \ k" + using assms(4,5) wf_item_def by force + ultimately have "derives \ [item_rule_head x] (slice i k \ @ tl (item_\ x))" + by (metis Derivation_implies_derives append.assoc slice_concat) + thus "sound_item \ \ (Item r\<^sub>x (b\<^sub>x + 1) i k)" + using assms(1,4) by (auto simp: sound_item_def item_defs drop_Suc tl_drop) +qed + +lemma sound_Earley: + assumes "x \ Earley \ \" "wf_item \ \ x" + shows "sound_item \ \ x" + using assms +proof (induction rule: Earley.induct) + case (Init r) + thus ?case + using sound_Init by blast +next + case (Scan x r b i j a) + thus ?case + using wf_Earley sound_Scan by fast +next + case (Predict x r b i j r') + thus ?case + using wf_Earley sound_Predict by blast +next + case (Complete x r\<^sub>x b\<^sub>x i j y r\<^sub>y b\<^sub>y k) + thus ?case + using wf_Earley sound_Complete by metis +qed + +theorem soundness_Earley: + assumes "recognizing (Earley \ \) \ \" + shows "derives \ [\ \] \" +proof - + obtain x where x: "x \ Earley \ \" "is_finished \ \ x" + using assms recognizing_def by blast + hence "sound_item \ \ x" + using wf_Earley sound_Earley by blast + thus ?thesis + unfolding sound_item_def using x by (auto simp: is_finished_def is_complete_def item_defs) +qed + + +subsection \Completeness\ + +definition partially_completed :: "nat \ 'a cfg \ 'a sentence \ 'a item set \ ('a derivation \ bool) \ bool" where + "partially_completed k \ \ E P \ \r b i' i j x a D. + i \ j \ j \ k \ k \ length \ \ + x = Item r b i' i \ x \ E \ next_symbol x = Some a \ + Derivation \ [a] D (slice i j \) \ P D \ + Item r (b+1) i' j \ E" + +lemma partially_completed_upto: + assumes "j \ k" "k \ length \" + assumes "x = Item (N,\) d i j" "x \ I" "\x \ I. wf_item \ \ x" + assumes "Derivation \ (item_\ x) D (slice j k \)" + assumes "partially_completed k \ \ I (\D'. length D' \ length D)" + shows "Item (N,\) (length \) i k \ I" + using assms +proof (induction "item_\ x" arbitrary: d i j k N \ x D) + case Nil + have "item_\ x = \" + using Nil(1,4) unfolding item_\_def item_\_def item_rule_body_def rule_body_def by simp + hence "x = Item (N,\) (length \) i j" + using Nil.hyps Nil.prems(3-5) unfolding wf_item_def item_defs by auto + have "Derivation \ [] D (slice j k \)" + using Nil.hyps Nil.prems(6) by auto + hence "slice j k \ = []" + using Derivation_from_empty by blast + hence "j = k" + unfolding slice_drop_take using Nil.prems(1,2) by simp + thus ?case + using \x = Item (N, \) (length \) i j\ Nil.prems(4) by blast +next + case (Cons b bs) + obtain j' E F where *: + "Derivation \ [b] E (slice j j' \)" + "Derivation \ bs F (slice j' k \)" + "j \ j'" "j' \ k" "length E \ length D" "length F \ length D" + using Derivation_concat_split[of \ "[b]" bs D "slice j k \"] slice_concat_Ex + using Cons.hyps(2) Cons.prems(1,6) + by (smt (verit, ccfv_threshold) Cons_eq_appendI append_self_conv2) + have "next_symbol x = Some b" + using Cons.hyps(2) unfolding item_defs(4) next_symbol_def is_complete_def by (auto, metis nth_via_drop) + hence "Item (N, \) (d+1) i j' \ I" + using Cons.prems(7) unfolding partially_completed_def + using Cons.prems(2,3,4) *(1,3-5) by blast + moreover have "partially_completed k \ \ I (\D'. length D' \ length F)" + using Cons.prems(7) *(6) unfolding partially_completed_def by fastforce + moreover have "bs = item_\ (Item (N,\) (d+1) i j')" + using Cons.hyps(2) Cons.prems(3) unfolding item_defs(4) item_rule_body_def + by (auto, metis List.list.sel(3) drop_Suc drop_tl) + ultimately show ?case + using Cons.hyps(1) *(2,4) Cons.prems(2,3,5) wf_item_def by blast +qed + +lemma partially_completed_Earley_k: + assumes "wf_\ \" + shows "partially_completed k \ \ (Earley \ \) (\_. True)" + unfolding partially_completed_def +proof (standard, standard, standard, standard, standard, standard, standard, standard, standard) + fix r b i' i j x a D + assume + "i \ j \ j \ k \ k \ length \ \ + x = Item r b i' i \ x \ Earley \ \ \ + next_symbol x = Some a \ + Derivation \ [a] D (slice i j \) \ True" + thus "Item r (b + 1) i' j \ Earley \ \" + proof (induction "length D" arbitrary: r b i' i j x a D rule: nat_less_induct) + case 1 + show ?case + proof cases + assume "D = []" + hence "[a] = slice i j \" + using "1.prems" by force + moreover have "j \ length \" + using le_trans "1.prems" by blast + ultimately have "j = i+1" + using slice_singleton by metis + hence "i < length \" + using \j \ length \\ discrete by blast + hence "\!i = a" + using slice_nth \[a] = slice i j \\ \j = i + 1\ by fastforce + hence "Item r (b + 1) i' j \ Earley \ \" + using Earley.Scan "1.prems" \i < length \\ \j = i + 1\ by metis + thus ?thesis + by (simp add: \j = i + 1\) + next + assume "\ D = []" + then obtain d D' where "D = d # D'" + by (meson List.list.exhaust) + then obtain \ where *: "Derives1 \ [a] (fst d) (snd d) \" "Derivation \ \ D' (slice i j \)" + using "1.prems" by auto + hence rule: "(a, \) \ set (\ \)" "fst d = 0" "snd d = (a ,\)" + using *(1) unfolding Derives1_def by (simp add: Cons_eq_append_conv)+ + show ?thesis + proof cases + assume "is_terminal \ a" + have "is_nonterminal \ a" + using rule by (simp add: assms) + thus ?thesis + using \is_terminal \ a\ is_terminal_nonterminal by (metis assms) + next + assume "\ is_terminal \ a" + define y where y_def: "y = Item (a ,\) 0 i i" + have "length D' < length D" + using \D = d # D'\ by fastforce + hence "partially_completed k \ \ (Earley \ \) (\E. length E \ length D')" + unfolding partially_completed_def using "1.hyps" order_le_less_trans by (smt (verit, best)) + hence "partially_completed j \ \ (Earley \ \) (\E. length E \ length D')" + unfolding partially_completed_def using "1.prems" by force + moreover have "Derivation \ (item_\ y) D' (slice i j \)" + using *(2) by (auto simp: item_defs y_def) + moreover have "y \ Earley \ \" + using y_def "1.prems" rule by (auto simp: item_defs Earley.Predict) + moreover have "j \ length \" + using "1.prems" by simp + ultimately have "Item (a,\) (length \) i j \ Earley \ \" + using partially_completed_upto "1.prems" wf_Earley y_def by metis + moreover have x: "x = Item r b i' i" "x \ Earley \ \" + using "1.prems" by blast+ + moreover have "next_symbol x = Some a" + using "1.prems" by linarith + ultimately show ?thesis + using Earley.Complete[OF x] by (auto simp: is_complete_def item_defs) + qed + qed + qed +qed + +lemma partially_completed_Earley: + "wf_\ \ \ partially_completed (length \) \ \ (Earley \ \) (\_. True)" + by (simp add: partially_completed_Earley_k) + +theorem completeness_Earley: + assumes "derives \ [\ \] \" "is_word \ \" "wf_\ \" + shows "recognizing (Earley \ \) \ \" +proof - + obtain \ D where *: "(\ \ ,\) \ set (\ \)" "Derivation \ \ D \" + using Derivation_\1 assms derives_implies_Derivation by metis + define x where x_def: "x = Item (\ \, \) 0 0 0" + have "partially_completed (length \) \ \ (Earley \ \) (\_. True)" + using assms(3) partially_completed_Earley by blast + hence 0: "partially_completed (length \) \ \ (Earley \ \) (\D'. length D' \ length D)" + unfolding partially_completed_def by blast + have 1: "x \ Earley \ \" + using x_def Earley.Init *(1) by fastforce + have 2: "Derivation \ (item_\ x) D (slice 0 (length \) \)" + using *(2) x_def by (simp add: item_defs) + have "Item (\ \,\) (length \) 0 (length \) \ Earley \ \" + using partially_completed_upto[OF _ _ _ _ _ 2 0] wf_Earley 1 x_def by auto + then show ?thesis + unfolding recognizing_def is_finished_def by (auto simp: is_complete_def item_defs, force) +qed + + +subsection \Correctness\ + +theorem correctness_Earley: + assumes "wf_\ \" "is_word \ \" + shows "recognizing (Earley \ \) \ \ \ derives \ [\ \] \" + using assms soundness_Earley completeness_Earley by blast + + +subsection \Finiteness\ + +lemma finiteness_empty: + "set (\ \) = {} \ finite { x | x. wf_item \ \ x }" + unfolding wf_item_def by simp + +fun item_intro :: "'a rule \ nat \ nat \ nat \ 'a item" where + "item_intro (rule, dot, origin, ends) = Item rule dot origin ends" + +lemma finiteness_nonempty: + assumes "set (\ \) \ {}" + shows "finite { x | x. wf_item \ \ x }" +proof - + define M where "M = Max { length (rule_body r) | r. r \ set (\ \) }" + define Top where "Top = (set (\ \) \ {0..M} \ {0..length \} \ {0..length \})" + hence "finite Top" + using finite_cartesian_product finite by blast + have "inj_on item_intro Top" + unfolding Top_def inj_on_def by simp + hence "finite (item_intro ` Top)" + using finite_image_iff \finite Top\ by auto + have "{ x | x. wf_item \ \ x } \ item_intro ` Top" + proof standard + fix x + assume "x \ { x | x. wf_item \ \ x }" + then obtain rule dot origin endp where *: "x = Item rule dot origin endp" + "rule \ set (\ \)" "dot \ length (item_rule_body x)" "origin \ length \" "endp \ length \" + unfolding wf_item_def using item.exhaust_sel le_trans by blast + hence "length (rule_body rule) \ { length (rule_body r) | r. r \ set (\ \) }" + using *(1,2) item_rule_body_def by blast + moreover have "finite { length (rule_body r) | r. r \ set (\ \) }" + using finite finite_image_set[of "\x. x \ set (\ \)"] by fastforce + ultimately have "M \ length (rule_body rule)" + unfolding M_def by simp + hence "dot \ M" + using *(1,3) item_rule_body_def by (metis item.sel(1) le_trans) + hence "(rule, dot, origin, endp) \ Top" + using *(2,4,5) unfolding Top_def by simp + thus "x \ item_intro ` Top" + using *(1) by force + qed + thus ?thesis + using \finite (item_intro ` Top)\ rev_finite_subset by auto +qed + +lemma finiteness_UNIV_wf_item: + "finite { x | x. wf_item \ \ x }" + using finiteness_empty finiteness_nonempty by fastforce + +theorem finiteness_Earley: + "finite (Earley \ \)" + using finiteness_UNIV_wf_item wf_Earley rev_finite_subset by (metis mem_Collect_eq subsetI) + +end \ No newline at end of file diff --git a/thys/Earley_Parser/Earley_Fixpoint.thy b/thys/Earley_Parser/Earley_Fixpoint.thy new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/Earley_Fixpoint.thy @@ -0,0 +1,508 @@ +theory Earley_Fixpoint + imports + Earley + Limit +begin + +section \Earley recognizer\ + +subsection \Earley fixpoint\ + +definition init_item :: "'a rule \ nat \ 'a item" where + "init_item r k \ Item r 0 k k" + +definition inc_item :: "'a item \ nat \ 'a item" where + "inc_item x k \ Item (item_rule x) (item_dot x + 1) (item_origin x) k" + +definition bin :: "'a item set \ nat \ 'a item set" where + "bin I k \ { x . x \ I \ item_end x = k }" + +definition Init\<^sub>F :: "'a cfg \ 'a item set" where + "Init\<^sub>F \ \ { init_item r 0 | r. r \ set (\ \) \ fst r = (\ \) }" + +definition Scan\<^sub>F :: "nat \ 'a sentence \ 'a item set \ 'a item set" where + "Scan\<^sub>F k \ I \ { inc_item x (k+1) | x a. + x \ bin I k \ + \!k = a \ + k < length \ \ + next_symbol x = Some a }" + +definition Predict\<^sub>F :: "nat \ 'a cfg \ 'a item set \ 'a item set" where + "Predict\<^sub>F k \ I \ { init_item r k | r x. + r \ set (\ \) \ + x \ bin I k \ + next_symbol x = Some (rule_head r) }" + +definition Complete\<^sub>F :: "nat \ 'a item set \ 'a item set" where + "Complete\<^sub>F k I \ { inc_item x k | x y. + x \ bin I (item_origin y) \ + y \ bin I k \ + is_complete y \ + next_symbol x = Some (item_rule_head y) }" + +definition Earley\<^sub>F_bin_step :: "nat \ 'a cfg \ 'a sentence \ 'a item set \ 'a item set" where + "Earley\<^sub>F_bin_step k \ \ I \ I \ Scan\<^sub>F k \ I \ Complete\<^sub>F k I \ Predict\<^sub>F k \ I" + +definition Earley\<^sub>F_bin :: "nat \ 'a cfg \ 'a sentence \ 'a item set \ 'a item set" where + "Earley\<^sub>F_bin k \ \ I \ limit (Earley\<^sub>F_bin_step k \ \) I" + +fun Earley\<^sub>F_bins :: "nat \ 'a cfg \ 'a sentence \ 'a item set" where + "Earley\<^sub>F_bins 0 \ \ = Earley\<^sub>F_bin 0 \ \ (Init\<^sub>F \)" +| "Earley\<^sub>F_bins (Suc n) \ \ = Earley\<^sub>F_bin (Suc n) \ \ (Earley\<^sub>F_bins n \ \)" + +definition Earley\<^sub>F :: "'a cfg \ 'a sentence \ 'a item set" where + "Earley\<^sub>F \ \ \ Earley\<^sub>F_bins (length \) \ \" + + +subsection \Monotonicity and Absorption\ + +lemma Earley\<^sub>F_bin_step_empty: + "Earley\<^sub>F_bin_step k \ \ {} = {}" + unfolding Earley\<^sub>F_bin_step_def Scan\<^sub>F_def Complete\<^sub>F_def Predict\<^sub>F_def bin_def by blast + +lemma Earley\<^sub>F_bin_step_setmonotone: + "setmonotone (Earley\<^sub>F_bin_step k \ \)" + by (simp add: Un_assoc Earley\<^sub>F_bin_step_def setmonotone_def) + +lemma Earley\<^sub>F_bin_step_continuous: + "continuous (Earley\<^sub>F_bin_step k \ \)" + unfolding continuous_def +proof (standard, standard, standard) + fix C :: "nat \ 'a item set" + assume "chain C" + thus "chain (Earley\<^sub>F_bin_step k \ \ \ C)" + unfolding chain_def Earley\<^sub>F_bin_step_def by (auto simp: Scan\<^sub>F_def Predict\<^sub>F_def Complete\<^sub>F_def bin_def subset_eq) +next + fix C :: "nat \ 'a item set" + assume *: "chain C" + show "Earley\<^sub>F_bin_step k \ \ (natUnion C) = natUnion (Earley\<^sub>F_bin_step k \ \ \ C)" + unfolding natUnion_def + proof standard + show "Earley\<^sub>F_bin_step k \ \ (\ {C n |n. True}) \ \ {(Earley\<^sub>F_bin_step k \ \ \ C) n |n. True}" + proof standard + fix x + assume #: "x \ Earley\<^sub>F_bin_step k \ \ (\ {C n |n. True})" + show "x \ \ {(Earley\<^sub>F_bin_step k \ \ \ C) n |n. True}" + proof (cases "x \ Complete\<^sub>F k (\ {C n |n. True})") + case True + then show ?thesis + using * unfolding chain_def Earley\<^sub>F_bin_step_def Complete\<^sub>F_def bin_def + proof clarsimp + fix y :: "'a item" and z :: "'a item" and n :: nat and m :: nat + assume a1: "is_complete z" + assume a2: "item_end y = item_origin z" + assume a3: "y \ C n" + assume a4: "z \ C m" + assume a5: "next_symbol y = Some (item_rule_head z)" + assume "\i. C i \ C (Suc i)" + hence f6: "\n m. \ n \ m \ C n \ C m" + by (meson lift_Suc_mono_le) + hence f7: "\n. \ m \ n \ z \ C n" + using a4 by blast + have "\n \ m. y \ C n" + using f6 a3 by (meson le_sup_iff subset_eq sup_ge1) + thus "\I. + (\n. I = C n \ + Scan\<^sub>F (item_end z) \ (C n) \ + {inc_item i (item_end z) |i. + i \ C n \ + (\j. + item_end i = item_origin j \ + j \ C n \ + item_end j = item_end z \ + is_complete j \ + next_symbol i = Some (item_rule_head j))} \ + Predict\<^sub>F (item_end z) \ (C n)) + \ inc_item y (item_end z) \ I" + using f7 a5 a2 a1 by blast + qed + next + case False + thus ?thesis + using # Un_iff by (auto simp: Earley\<^sub>F_bin_step_def Scan\<^sub>F_def Predict\<^sub>F_def bin_def; blast) + qed + qed + next + show "\ {(Earley\<^sub>F_bin_step k \ \ \ C) n |n. True} \ Earley\<^sub>F_bin_step k \ \ (\ {C n |n. True})" + unfolding Earley\<^sub>F_bin_step_def + using * by (auto simp: Scan\<^sub>F_def Predict\<^sub>F_def Complete\<^sub>F_def chain_def bin_def, metis+) + qed +qed + +lemma Earley\<^sub>F_bin_step_regular: + "regular (Earley\<^sub>F_bin_step k \ \)" + by (simp add: Earley\<^sub>F_bin_step_continuous Earley\<^sub>F_bin_step_setmonotone regular_def) + +lemma Earley\<^sub>F_bin_idem: + "Earley\<^sub>F_bin k \ \ (Earley\<^sub>F_bin k \ \ I) = Earley\<^sub>F_bin k \ \ I" + by (simp add: Earley\<^sub>F_bin_def Earley\<^sub>F_bin_step_regular limit_is_idempotent) + +lemma Scan\<^sub>F_bin_absorb: + "Scan\<^sub>F k \ (bin I k) = Scan\<^sub>F k \ I" + unfolding Scan\<^sub>F_def bin_def by simp + +lemma Predict\<^sub>F_bin_absorb: + "Predict\<^sub>F k \ (bin I k) = Predict\<^sub>F k \ I" + unfolding Predict\<^sub>F_def bin_def by simp + +lemma Scan\<^sub>F_Un: + "Scan\<^sub>F k \ (I \ J) = Scan\<^sub>F k \ I \ Scan\<^sub>F k \ J" + unfolding Scan\<^sub>F_def bin_def by blast + +lemma Predict\<^sub>F_Un: + "Predict\<^sub>F k \ (I \ J) = Predict\<^sub>F k \ I \ Predict\<^sub>F k \ J" + unfolding Predict\<^sub>F_def bin_def by blast + +lemma Scan\<^sub>F_sub_mono: + "I \ J \ Scan\<^sub>F k \ I \ Scan\<^sub>F k \ J" + unfolding Scan\<^sub>F_def bin_def by blast + +lemma Predict\<^sub>F_sub_mono: + "I \ J \ Predict\<^sub>F k \ I \ Predict\<^sub>F k \ J" + unfolding Predict\<^sub>F_def bin_def by blast + +lemma Complete\<^sub>F_sub_mono: + "I \ J \ Complete\<^sub>F k I \ Complete\<^sub>F k J" + unfolding Complete\<^sub>F_def bin_def by blast + +lemma Earley\<^sub>F_bin_step_sub_mono: + "I \ J \ Earley\<^sub>F_bin_step k \ \ I \ Earley\<^sub>F_bin_step k \ \ J" + unfolding Earley\<^sub>F_bin_step_def using Scan\<^sub>F_sub_mono Predict\<^sub>F_sub_mono Complete\<^sub>F_sub_mono by (metis sup.mono) + +lemma funpower_sub_mono: + "I \ J \ funpower (Earley\<^sub>F_bin_step k \ \) n I \ funpower (Earley\<^sub>F_bin_step k \ \) n J" + by (induction n) (auto simp: Earley\<^sub>F_bin_step_sub_mono) + +lemma Earley\<^sub>F_bin_sub_mono: + "I \ J \ Earley\<^sub>F_bin k \ \ I \ Earley\<^sub>F_bin k \ \ J" +proof standard + fix x + assume "I \ J" "x \ Earley\<^sub>F_bin k \ \ I" + then obtain n where "x \ funpower (Earley\<^sub>F_bin_step k \ \) n I" + unfolding Earley\<^sub>F_bin_def limit_def natUnion_def by blast + hence "x \ funpower (Earley\<^sub>F_bin_step k \ \) n J" + using \I \ J\ funpower_sub_mono by blast + thus "x \ Earley\<^sub>F_bin k \ \ J" + unfolding Earley\<^sub>F_bin_def limit_def natUnion_def by blast +qed + +lemma Scan\<^sub>F_Earley\<^sub>F_bin_step_mono: + "Scan\<^sub>F k \ I \ Earley\<^sub>F_bin_step k \ \ I" + using Earley\<^sub>F_bin_step_def by blast + +lemma Predict\<^sub>F_Earley\<^sub>F_bin_step_mono: + "Predict\<^sub>F k \ I \ Earley\<^sub>F_bin_step k \ \ I" + using Earley\<^sub>F_bin_step_def by blast + +lemma Complete\<^sub>F_Earley\<^sub>F_bin_step_mono: + "Complete\<^sub>F k I \ Earley\<^sub>F_bin_step k \ \ I" + using Earley\<^sub>F_bin_step_def by blast + +lemma Earley\<^sub>F_bin_step_Earley\<^sub>F_bin_mono: + "Earley\<^sub>F_bin_step k \ \ I \ Earley\<^sub>F_bin k \ \ I" +proof - + have "Earley\<^sub>F_bin_step k \ \ I \ funpower (Earley\<^sub>F_bin_step k \ \) 1 I" + by simp + thus ?thesis + by (metis Earley\<^sub>F_bin_def limit_elem subset_eq) +qed + +lemma Scan\<^sub>F_Earley\<^sub>F_bin_mono: + "Scan\<^sub>F k \ I \ Earley\<^sub>F_bin k \ \ I" + using Scan\<^sub>F_Earley\<^sub>F_bin_step_mono Earley\<^sub>F_bin_step_Earley\<^sub>F_bin_mono by force + +lemma Predict\<^sub>F_Earley\<^sub>F_bin_mono: + "Predict\<^sub>F k \ I \ Earley\<^sub>F_bin k \ \ I" + using Predict\<^sub>F_Earley\<^sub>F_bin_step_mono Earley\<^sub>F_bin_step_Earley\<^sub>F_bin_mono by force + +lemma Complete\<^sub>F_Earley\<^sub>F_bin_mono: + "Complete\<^sub>F k I \ Earley\<^sub>F_bin k \ \ I" + using Complete\<^sub>F_Earley\<^sub>F_bin_step_mono Earley\<^sub>F_bin_step_Earley\<^sub>F_bin_mono by force + +lemma Earley\<^sub>F_bin_mono: + "I \ Earley\<^sub>F_bin k \ \ I" + using Earley\<^sub>F_bin_step_Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_step_def by blast + +lemma Init\<^sub>F_sub_Earley\<^sub>F_bins: + "Init\<^sub>F \ \ Earley\<^sub>F_bins n \ \" + by (induction n) (use Earley\<^sub>F_bin_mono in fastforce)+ + +subsection \Soundness\ + +lemma Init\<^sub>F_sub_Earley: + "Init\<^sub>F \ \ Earley \ \" + unfolding Init\<^sub>F_def init_item_def using Init by blast + +lemma Scan\<^sub>F_sub_Earley: + assumes "I \ Earley \ \" + shows "Scan\<^sub>F k \ I \ Earley \ \" + unfolding Scan\<^sub>F_def inc_item_def bin_def using assms Scan + by (smt (verit, ccfv_SIG) item.exhaust_sel mem_Collect_eq subsetD subsetI) + +lemma Predict\<^sub>F_sub_Earley: + assumes "I \ Earley \ \" + shows "Predict\<^sub>F k \ I \ Earley \ \" + unfolding Predict\<^sub>F_def init_item_def bin_def using assms Predict + using item.exhaust_sel by blast + +lemma Complete\<^sub>F_sub_Earley: + assumes "I \ Earley \ \" + shows "Complete\<^sub>F k I \ Earley \ \" + unfolding Complete\<^sub>F_def inc_item_def bin_def using assms Complete + by (smt (verit, del_insts) item.exhaust_sel mem_Collect_eq subset_eq) + +lemma Earley\<^sub>F_bin_step_sub_Earley: + assumes "I \ Earley \ \" + shows "Earley\<^sub>F_bin_step k \ \ I \ Earley \ \" + unfolding Earley\<^sub>F_bin_step_def using assms Complete\<^sub>F_sub_Earley Predict\<^sub>F_sub_Earley Scan\<^sub>F_sub_Earley by (metis le_supI) + +lemma Earley\<^sub>F_bin_sub_Earley: + assumes "I \ Earley \ \" + shows "Earley\<^sub>F_bin k \ \ I \ Earley \ \" + using assms Earley\<^sub>F_bin_step_sub_Earley by (metis Earley\<^sub>F_bin_def limit_upperbound) + +lemma Earley\<^sub>F_bins_sub_Earley: + shows "Earley\<^sub>F_bins n \ \ \ Earley \ \" + by (induction n) (auto simp: Earley\<^sub>F_bin_sub_Earley Init\<^sub>F_sub_Earley) + +lemma Earley\<^sub>F_sub_Earley: + shows "Earley\<^sub>F \ \ \ Earley \ \" + by (simp add: Earley\<^sub>F_bins_sub_Earley Earley\<^sub>F_def) + +theorem soundness_Earley\<^sub>F: + assumes "recognizing (Earley\<^sub>F \ \) \ \" + shows "derives \ [\ \] \" + using soundness_Earley Earley\<^sub>F_sub_Earley assms recognizing_def by (metis subsetD) + + +subsection \Completeness\ + +definition prev_symbol :: "'a item \ 'a option" where + "prev_symbol x \ if item_dot x = 0 then None else Some (item_rule_body x ! (item_dot x - 1))" + +definition base :: "'a sentence \ 'a item set \ nat \ 'a item set" where + "base \ I k \ { x . x \ I \ item_end x = k \ k > 0 \ prev_symbol x = Some (\!(k-1)) }" + +lemma Earley\<^sub>F_bin_sub_Earley\<^sub>F_bin: + assumes "Init\<^sub>F \ \ I" + assumes "\k' < k. bin (Earley \ \) k' \ I" + assumes "base \ (Earley \ \) k \ I" + shows "bin (Earley \ \) k \ bin (Earley\<^sub>F_bin k \ \ I) k" +proof standard + fix x + assume *: "x \ bin (Earley \ \) k" + hence "x \ Earley \ \" + using bin_def by blast + thus "x \ bin (Earley\<^sub>F_bin k \ \ I) k" + using assms * + proof (induction rule: Earley.induct) + case (Init r) + thus ?case + unfolding Init\<^sub>F_def init_item_def bin_def using Earley\<^sub>F_bin_mono by fast + next + case (Scan x r b i j a) + have "j+1 = k" + using Scan.prems(4) bin_def by (metis (mono_tags, lifting) CollectD item.sel(4)) + have "prev_symbol (Item r (b+1) i (j+1)) = Some (\!(k-1))" + using Scan.hyps(1,3,5) \j+1 = k\ by (auto simp: next_symbol_def prev_symbol_def item_rule_body_def split: if_splits) + hence "Item r (b+1) i (j+1) \ base \ (Earley \ \) k" + unfolding base_def using Scan.prems(4) bin_def by fastforce + hence "Item r (b+1) i (j+1) \ I" + using Scan.prems(3) by blast + hence "Item r (b+1) i (j+1) \ Earley\<^sub>F_bin k \ \ I" + using Earley\<^sub>F_bin_mono by blast + thus ?case + using \j+1 = k\ bin_def by fastforce + next + case (Predict x r b i j r') + have "j = k" + using Predict.prems(4) bin_def by (metis (mono_tags, lifting) CollectD item.sel(4)) + hence "x \ bin (Earley \ \) k" + using Predict.hyps(1,2) bin_def by fastforce + hence "x \ bin (Earley\<^sub>F_bin k \ \ I) k" + using Predict.IH Predict.prems(1-3) by blast + hence "Item r' 0 j j \ Predict\<^sub>F k \ (Earley\<^sub>F_bin k \ \ I)" + unfolding Predict\<^sub>F_def init_item_def using Predict.hyps(1,3,4) \j = k\ by blast + hence "Item r' 0 j j \ Earley\<^sub>F_bin_step k \ \ (Earley\<^sub>F_bin k \ \ I)" + using Predict\<^sub>F_Earley\<^sub>F_bin_step_mono by blast + hence "Item r' 0 j j \ Earley\<^sub>F_bin k \ \ I" + using Earley\<^sub>F_bin_idem Earley\<^sub>F_bin_step_Earley\<^sub>F_bin_mono by blast + thus ?case + by (simp add: \j = k\ bin_def) + next + case (Complete x r\<^sub>x b\<^sub>x i j y r\<^sub>y b\<^sub>y l) + have "l = k" + using Complete.prems(4) bin_def by (metis (mono_tags, lifting) CollectD item.sel(4)) + hence "y \ bin (Earley \ \) l" + using Complete.hyps(3,4) bin_def by fastforce + hence 0: "y \ bin (Earley\<^sub>F_bin k \ \ I) k" + using Complete.IH(2) Complete.prems(1-3) \l = k\ by blast + have 1: "x \ bin (Earley\<^sub>F_bin k \ \ I) (item_origin y)" + proof (cases "j = k") + case True + hence "x \ bin (Earley \ \) k" + using Complete.hyps(1,2) bin_def by fastforce + hence "x \ bin (Earley\<^sub>F_bin k \ \ I) k" + using Complete.IH(1) Complete.prems(1-3) by blast + thus ?thesis + using Complete.hyps(3) True by simp + next + case False + hence "j < k" + using \l = k\ wf_Earley wf_item_def Complete.hyps(3,4) by force + moreover have "x \ bin (Earley \ \) j" + using Complete.hyps(1,2) bin_def by force + ultimately have "x \ I" + using Complete.prems(2) by blast + hence "x \ bin (Earley\<^sub>F_bin k \ \ I) j" + using Complete.hyps(1) Earley\<^sub>F_bin_mono bin_def by fastforce + thus ?thesis + using Complete.hyps(3) by simp + qed + have "Item r\<^sub>x (b\<^sub>x + 1) i k \ Complete\<^sub>F k (Earley\<^sub>F_bin k \ \ I)" + unfolding Complete\<^sub>F_def inc_item_def using 0 1 Complete.hyps(1,5,6) by force + hence "Item r\<^sub>x (b\<^sub>x + 1) i k \ Earley\<^sub>F_bin_step k \ \ (Earley\<^sub>F_bin k \ \ I)" + unfolding Earley\<^sub>F_bin_step_def by blast + hence "Item r\<^sub>x (b\<^sub>x + 1) i k \ Earley\<^sub>F_bin k \ \ I" + using Earley\<^sub>F_bin_idem Earley\<^sub>F_bin_step_Earley\<^sub>F_bin_mono by blast + thus ?case + using bin_def \l = k\ by fastforce + qed +qed + +lemma Earley_base_sub_Earley\<^sub>F_bin: + assumes "Init\<^sub>F \ \ I" + assumes "\k' < k. bin (Earley \ \) k' \ I" + assumes "base \ (Earley \ \) k \ I" + assumes "wf_\ \" "is_word \ \" + shows "base \ (Earley \ \) (k+1) \ bin (Earley\<^sub>F_bin k \ \ I) (k+1)" +proof standard + fix x + assume *: "x \ base \ (Earley \ \) (k+1)" + hence "x \ Earley \ \" + using base_def by blast + thus "x \ bin (Earley\<^sub>F_bin k \ \ I) (k+1)" + using assms * + proof (induction rule: Earley.induct) + case (Init r) + have "k = 0" + using Init.prems(6) unfolding base_def by simp + hence False + using Init.prems(6) unfolding base_def by simp + thus ?case + by blast + next + case (Scan x r b i j a) + have "j = k" + using Scan.prems(6) base_def by (metis (mono_tags, lifting) CollectD add_right_cancel item.sel(4)) + hence "x \ bin (Earley\<^sub>F_bin k \ \ I) k" + using Earley\<^sub>F_bin_sub_Earley\<^sub>F_bin Scan.prems Scan.hyps(1,2) bin_def + by (metis (mono_tags, lifting) CollectI item.sel(4) subsetD) + hence "Item r (b+1) i (j+1) \ Scan\<^sub>F k \ (Earley\<^sub>F_bin k \ \ I)" + unfolding Scan\<^sub>F_def inc_item_def using Scan.hyps \j = k\ by force + hence "Item r (b+1) i (j+1) \ Earley\<^sub>F_bin_step k \ \ (Earley\<^sub>F_bin k \ \ I)" + using Scan\<^sub>F_Earley\<^sub>F_bin_step_mono by blast + hence "Item r (b+1) i (j+1) \ Earley\<^sub>F_bin k \ \ I" + using Earley\<^sub>F_bin_idem Earley\<^sub>F_bin_step_Earley\<^sub>F_bin_mono by blast + thus ?case + using \j = k\ bin_def by fastforce + next + case (Predict x r b i j r') + have False + using Predict.prems(6) unfolding base_def by (auto simp: prev_symbol_def) + thus ?case + by blast + next + case (Complete x r\<^sub>x b\<^sub>x i j y r\<^sub>y b\<^sub>y l) + have "l-1 < length \" + using Complete.prems(6) base_def wf_Earley wf_item_def + by (metis (mono_tags, lifting) CollectD add.right_neutral add_Suc_right add_diff_cancel_right' item.sel(4) less_eq_Suc_le plus_1_eq_Suc) + hence "is_terminal \ (\!(l-1))" + using Complete.prems(5) is_word_is_terminal by blast + moreover have "is_nonterminal \ (item_rule_head y)" + using Complete.hyps(3,4) Complete.prems(4) wf_Earley wf_item_def + by (metis item_rule_head_def prod.collapse rule_head_def rule_nonterminal_type) + moreover have "prev_symbol (Item r\<^sub>x (b\<^sub>x+1) i l) = next_symbol x" + using Complete.hyps(1,6) + by (auto simp: next_symbol_def prev_symbol_def is_complete_def item_rule_body_def split: if_splits) + moreover have "prev_symbol (Item r\<^sub>x (b\<^sub>x+1) i l) = Some (\!(l-1))" + using Complete.prems(6) base_def by (metis (mono_tags, lifting) CollectD item.sel(4)) + ultimately have False + using Complete.hyps(6) Complete.prems(4) is_terminal_nonterminal by fastforce + thus ?case + by blast + qed +qed + +lemma Earley\<^sub>F_bin_k_sub_Earley\<^sub>F_bins: + assumes "wf_\ \" "is_word \ \" "k \ n" + shows "bin (Earley \ \) k \ Earley\<^sub>F_bins n \ \" + using assms +proof (induction n arbitrary: k) + case 0 + have "bin (Earley \ \) 0 \ bin (Earley\<^sub>F_bin 0 \ \ (Init\<^sub>F \)) 0" + using Earley\<^sub>F_bin_sub_Earley\<^sub>F_bin base_def by fastforce + thus ?case + unfolding bin_def using "0.prems"(3) by auto +next + case (Suc n) + show ?case + proof (cases "k \ n") + case True + thus ?thesis + using Suc Earley\<^sub>F_bin_mono by force + next + case False + hence "k = n+1" + using Suc.prems(3) by force + have 0: "\k' < k. bin (Earley \ \) k' \ Earley\<^sub>F_bins n \ \" + using Suc by simp + moreover have "base \ (Earley \ \) k \ Earley\<^sub>F_bins n \ \" + proof - + have "\k' < k-1. bin (Earley \ \) k' \ Earley\<^sub>F_bins n \ \" + using Suc \k = n + 1\ by auto + moreover have "base \ (Earley \ \) (k-1) \ Earley\<^sub>F_bins n \ \" + using 0 bin_def base_def False \k = n+1\ + by (smt (verit) Suc_eq_plus1 diff_Suc_1 linorder_not_less mem_Collect_eq subsetD subsetI) + ultimately have "base \ (Earley \ \) k \ bin (Earley\<^sub>F_bin n \ \ (Earley\<^sub>F_bins n \ \)) k" + using Suc.prems(1,2) Earley_base_sub_Earley\<^sub>F_bin \k = n + 1\ Init\<^sub>F_sub_Earley\<^sub>F_bins by (metis add_diff_cancel_right') + hence "base \ (Earley \ \) k \ bin (Earley\<^sub>F_bins n \ \) k" + by (metis Earley\<^sub>F_bins.elims Earley\<^sub>F_bin_idem) + thus ?thesis + using bin_def by blast + qed + ultimately have "bin (Earley \ \) k \ bin (Earley\<^sub>F_bin k \ \ (Earley\<^sub>F_bins n \ \)) k" + using Earley\<^sub>F_bin_sub_Earley\<^sub>F_bin Init\<^sub>F_sub_Earley\<^sub>F_bins by metis + thus ?thesis + using Earley\<^sub>F_bins.simps(2) \k = n + 1\ bin_def by auto + qed +qed + +lemma Earley_sub_Earley\<^sub>F: + assumes "wf_\ \" "is_word \ \" + shows "Earley \ \ \ Earley\<^sub>F \ \" +proof - + have "\k \ length \. bin (Earley \ \) k \ Earley\<^sub>F \ \" + by (simp add: Earley\<^sub>F_bin_k_sub_Earley\<^sub>F_bins Earley\<^sub>F_def assms) + thus ?thesis + using wf_Earley wf_item_def bin_def by blast +qed + +theorem completeness_Earley\<^sub>F: + assumes "derives \ [\ \] \" "is_word \ \" "wf_\ \" + shows "recognizing (Earley\<^sub>F \ \) \ \" + using assms Earley_sub_Earley\<^sub>F Earley\<^sub>F_sub_Earley completeness_Earley by (metis subset_antisym) + +subsection \Correctness\ + +theorem Earley_eq_Earley\<^sub>F: + assumes "wf_\ \" "is_word \ \" + shows "Earley \ \ = Earley\<^sub>F \ \" + using Earley_sub_Earley\<^sub>F Earley\<^sub>F_sub_Earley assms by blast + +theorem correctness_Earley\<^sub>F: + assumes "wf_\ \" "is_word \ \" + shows "recognizing (Earley\<^sub>F \ \) \ \ \ derives \ [\ \] \" + using assms Earley_eq_Earley\<^sub>F correctness_Earley by fastforce + +end diff --git a/thys/Earley_Parser/Earley_Parser.thy b/thys/Earley_Parser/Earley_Parser.thy new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/Earley_Parser.thy @@ -0,0 +1,2528 @@ +theory Earley_Parser + imports + Earley_Recognizer + "HOL-Library.Monad_Syntax" +begin + +section \Earley parser\ + +subsection \Pointer lemmas\ + +definition predicts :: "'a item \ bool" where + "predicts x \ item_origin x = item_end x \ item_dot x = 0" + +definition scans :: "'a sentence \ nat \ 'a item \ 'a item \ bool" where + "scans \ k x y \ y = inc_item x k \ (\a. next_symbol x = Some a \ \!(k-1) = a)" + +definition completes :: "nat \ 'a item \ 'a item \ 'a item \ bool" where + "completes k x y z \ y = inc_item x k \ is_complete z \ item_origin z = item_end x \ + (\N. next_symbol x = Some N \ N = item_rule_head z)" + +definition sound_null_ptr :: "'a entry \ bool" where + "sound_null_ptr e \ (pointer e = Null \ predicts (item e))" + +definition sound_pre_ptr :: "'a sentence \ 'a bins \ nat \ 'a entry \ bool" where + "sound_pre_ptr \ bs k e \ \pre. pointer e = Pre pre \ + k > 0 \ pre < length (bs!(k-1)) \ scans \ k (item (bs!(k-1)!pre)) (item e)" + +definition sound_prered_ptr :: "'a bins \ nat \ 'a entry \ bool" where + "sound_prered_ptr bs k e \ \p ps k' pre red. pointer e = PreRed p ps \ (k', pre, red) \ set (p#ps) \ + k' < k \ pre < length (bs!k') \ red < length (bs!k) \ completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" + +definition sound_ptrs :: "'a sentence \ 'a bins \ bool" where + "sound_ptrs \ bs \ \k < length bs. \e \ set (bs!k). + sound_null_ptr e \ sound_pre_ptr \ bs k e \ sound_prered_ptr bs k e" + +definition mono_red_ptr :: "'a bins \ bool" where + "mono_red_ptr bs \ \k < length bs. \i < length (bs!k). + \k' pre red ps. pointer (bs!k!i) = PreRed (k', pre, red) ps \ red < i" + +lemma nth_item_bin_upd: + "n < length es \ item (bin_upd e es ! n) = item (es!n)" + by (induction es arbitrary: e n) (auto simp: less_Suc_eq_0_disj split: entry.splits pointer.splits) + +lemma bin_upd_append: + "item e \ set (items es) \ bin_upd e es = es @ [e]" + by (induction es arbitrary: e) (auto simp: items_def split: entry.splits pointer.splits) + +lemma bin_upd_null_pre: + "item e \ set (items es) \ pointer e = Null \ pointer e = Pre pre \ bin_upd e es = es" + by (induction es arbitrary: e) (auto simp: items_def split: entry.splits) + +lemma bin_upd_prered_nop: + assumes "distinct (items es)" "i < length es" + assumes "item e = item (es!i)" "pointer e = PreRed p ps" "\p ps. pointer (es!i) = PreRed p ps" + shows "bin_upd e es = es" + using assms + by (induction es arbitrary: e i) (auto simp: less_Suc_eq_0_disj items_def split: entry.splits pointer.splits) + +lemma bin_upd_prered_upd: + assumes "distinct (items es)" "i < length es" + assumes "item e = item (es!i)" "pointer e = PreRed p rs" "pointer (es!i) = PreRed p' rs'" "bin_upd e es = es'" + shows "pointer (es'!i) = PreRed p' (p#rs@rs') \ (\j < length es'. i\j \ es'!j = es!j) \ length (bin_upd e es) = length es" + using assms +proof (induction es arbitrary: e i es') + case (Cons e' es) + show ?case + proof cases + assume *: "item e = item e'" + show ?thesis + proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ e' = Entry y (PreRed yp ys)") + case True + then obtain x xp xs y yp ys where ee': "e = Entry x (PreRed xp xs)" "e' = Entry y (PreRed yp ys)" "x = y" + using * by auto + have simp: "bin_upd e (e' # es') = Entry x (PreRed yp (xp # xs @ ys)) # es'" + using True ee' by simp + show ?thesis + using Cons simp ee' apply (auto simp: items_def) + using less_Suc_eq_0_disj by fastforce+ + next + case False + hence "bin_upd e (e' # es') = e' # es'" + using * by (auto split: pointer.splits entry.splits) + thus ?thesis + using False * Cons.prems(1,2,3,4,5) by (auto simp: less_Suc_eq_0_disj items_def split: entry.splits) + qed + next + assume *: "item e \ item e'" + have simp: "bin_upd e (e' # es) = e' # bin_upd e es" + using * by (auto split: pointer.splits entry.splits) + have 0: "distinct (items es)" + using Cons.prems(1) unfolding items_def by simp + have 1: "i-1 < length es" + using Cons.prems(2,3) * by (metis One_nat_def leI less_diff_conv2 less_one list.size(4) nth_Cons_0) + have 2: "item e = item (es!(i-1))" + using Cons.prems(3) * by (metis nth_Cons') + have 3: "pointer e = PreRed p rs" + using Cons.prems(4) by simp + have 4: "pointer (es!(i-1)) = PreRed p' rs' " + using Cons.prems(3,5) * by (metis nth_Cons') + have "pointer (bin_upd e es!(i-1)) = PreRed p' (p # rs @ rs') \ + (\j < length (bin_upd e es). i-1 \ j \ (bin_upd e es) ! j = es ! j)" + using Cons.IH[OF 0 1 2 3 4] by blast + hence "pointer ((e' # bin_upd e es) ! i) = PreRed p' (p # rs @ rs') \ + (\j < length (e' # bin_upd e es). i \ j \ (e' # bin_upd e es) ! j = (e' # es) ! j)" + using * Cons.prems(2,3) less_Suc_eq_0_disj by auto + moreover have "e' # bin_upd e es = es'" + using Cons.prems(6) simp by auto + ultimately show ?thesis + by (metis 0 1 2 3 4 Cons.IH Cons.prems(6) length_Cons) + qed +qed simp + +lemma sound_ptrs_bin_upd: + assumes "sound_ptrs \ bs" "k < length bs" "es = bs!k" "distinct (items es)" + assumes "sound_null_ptr e" "sound_pre_ptr \ bs k e" "sound_prered_ptr bs k e" + shows "sound_ptrs \ (bs[k := bin_upd e es])" + unfolding sound_ptrs_def +proof (standard, standard, standard) + fix idx elem + let ?bs = "bs[k := bin_upd e es]" + assume a0: "idx < length ?bs" + assume a1: "elem \ set (?bs ! idx)" + show "sound_null_ptr elem \ sound_pre_ptr \ ?bs idx elem \ sound_prered_ptr ?bs idx elem" + proof cases + assume a2: "idx = k" + have "elem \ set es \ sound_pre_ptr \ bs idx elem" + using a0 a2 assms(1-3) sound_ptrs_def by blast + hence pre_es: "elem \ set es \ sound_pre_ptr \ ?bs idx elem" + using a2 unfolding sound_pre_ptr_def by force + have "elem = e \ sound_pre_ptr \ bs idx elem" + using a2 assms(6) by auto + hence pre_e: "elem = e \ sound_pre_ptr \ ?bs idx elem" + using a2 unfolding sound_pre_ptr_def by force + have "elem \ set es \ sound_prered_ptr bs idx elem" + using a0 a2 assms(1-3) sound_ptrs_def by blast + hence prered_es: "elem \ set es \ sound_prered_ptr (bs[k := bin_upd e es]) idx elem" + using a2 assms(2,3) length_bin_upd nth_item_bin_upd unfolding sound_prered_ptr_def + by (smt (verit, ccfv_SIG) dual_order.strict_trans1 nth_list_update) + have "elem = e \ sound_prered_ptr bs idx elem" + using a2 assms(7) by auto + hence prered_e: "elem = e \ sound_prered_ptr ?bs idx elem" + using a2 assms(2,3) length_bin_upd nth_item_bin_upd unfolding sound_prered_ptr_def + by (smt (verit, best) dual_order.strict_trans1 nth_list_update) + consider (A) "item e \ set (items es)" | + (B) "item e \ set (items es) \ (\pre. pointer e = Null \ pointer e = Pre pre)" | + (C) "item e \ set (items es) \ \ (\pre. pointer e = Null \ pointer e = Pre pre)" + by blast + thus ?thesis + proof cases + case A + hence "elem \ set (es @ [e])" + using a1 a2 bin_upd_append assms(2) by force + thus ?thesis + using assms(1-3,5) pre_e pre_es prered_e prered_es sound_ptrs_def by auto + next + case B + hence "elem \ set es" + using a1 a2 bin_upd_null_pre assms(2) by force + thus ?thesis + using assms(1-3) pre_es prered_es sound_ptrs_def by blast + next + case C + then obtain i p ps where C: "i < length es \ item e = item (es!i) \ pointer e = PreRed p ps" + by (metis assms(4) distinct_Ex1 items_def length_map nth_map pointer.exhaust) + show ?thesis + proof cases + assume "\p' ps'. pointer (es!i) = PreRed p' ps'" + hence C: "elem \ set es" + using a1 a2 C bin_upd_prered_nop assms(2,4) by (metis nth_list_update_eq) + thus ?thesis + using assms(1-3) sound_ptrs_def pre_es prered_es by blast + next + assume "\ (\p' ps'. pointer (es!i) = PreRed p' ps')" + then obtain p' ps' where D: "pointer (es!i) = PreRed p' ps'" + by blast + hence 0: "pointer (bin_upd e es!i) = PreRed p' (p#ps@ps') \ (\j < length (bin_upd e es). i\j \ bin_upd e es!j = es!j)" + using C assms(4) bin_upd_prered_upd by blast + obtain j where 1: "j < length es \ elem = bin_upd e es!j" + using a1 a2 assms(2) C items_def bin_eq_items_bin_upd by (metis in_set_conv_nth length_map nth_list_update_eq nth_map) + show ?thesis + proof cases + assume a3: "i=j" + hence a3: "pointer elem = PreRed p' (p#ps@ps')" + using 0 1 by blast + have "sound_null_ptr elem" + using a3 unfolding sound_null_ptr_def by simp + moreover have "sound_pre_ptr \ ?bs idx elem" + using a3 unfolding sound_pre_ptr_def by simp + moreover have "sound_prered_ptr ?bs idx elem" + unfolding sound_prered_ptr_def + proof (standard, standard, standard, standard, standard, standard) + fix P PS k' pre red + assume a4: "pointer elem = PreRed P PS \ (k', pre, red) \ set (P#PS)" + show "k' < idx \ pre < length (bs[k := bin_upd e es]!k') \ red < length (bs[k := bin_upd e es]!idx) \ + completes idx (item (bs[k := bin_upd e es]!k'!pre)) (item elem) (item (bs[k := bin_upd e es]!idx!red))" + proof cases + assume a5: "(k', pre, red) \ set (p#ps)" + show ?thesis + using 0 1 C a2 a4 a5 prered_es assms(2,3,7) sound_prered_ptr_def length_bin_upd nth_item_bin_upd + by (smt (verit) dual_order.strict_trans1 nth_list_update_eq nth_list_update_neq nth_mem) + next + assume a5: "(k', pre, red) \ set (p#ps)" + hence a5: "(k', pre, red) \ set (p'#ps')" + using a3 a4 by auto + have "k' < idx \ pre < length (bs!k') \ red < length (bs!idx) \ + completes idx (item (bs!k'!pre)) (item e) (item (bs!idx!red))" + using assms(1-3) C D a2 a5 unfolding sound_ptrs_def sound_prered_ptr_def by (metis nth_mem) + thus ?thesis + using 0 1 C a4 assms(2,3) length_bin_upd nth_item_bin_upd prered_es sound_prered_ptr_def + by (smt (verit, best) dual_order.strict_trans1 nth_list_update_eq nth_list_update_neq nth_mem) + qed + qed + ultimately show ?thesis + by blast + next + assume a3: "i\j" + hence "elem \ set es" + using 0 1 by (metis length_bin_upd nth_mem order_less_le_trans) + thus ?thesis + using assms(1-3) pre_es prered_es sound_ptrs_def by blast + qed + qed + qed + next + assume a2: "idx \ k" + have null: "sound_null_ptr elem" + using a0 a1 a2 assms(1) sound_ptrs_def by auto + have "sound_pre_ptr \ bs idx elem" + using a0 a1 a2 assms(1,2) unfolding sound_ptrs_def by simp + hence pre: "sound_pre_ptr \ ?bs idx elem" + using assms(2,3) length_bin_upd nth_item_bin_upd unfolding sound_pre_ptr_def + using dual_order.strict_trans1 nth_list_update by fastforce + have "sound_prered_ptr bs idx elem" + using a0 a1 a2 assms(1,2) unfolding sound_ptrs_def by simp + hence prered: "sound_prered_ptr ?bs idx elem" + using assms(2,3) length_bin_upd nth_item_bin_upd unfolding sound_prered_ptr_def + by (smt (verit, best) dual_order.strict_trans1 nth_list_update) + show ?thesis + using null pre prered by blast + qed +qed + +lemma mono_red_ptr_bin_upd: + assumes "mono_red_ptr bs" "k < length bs" "es = bs!k" "distinct (items es)" + assumes "\k' pre red ps. pointer e = PreRed (k', pre, red) ps \ red < length es" + shows "mono_red_ptr (bs[k := bin_upd e es])" + unfolding mono_red_ptr_def +proof (standard, standard) + fix idx + let ?bs = "bs[k := bin_upd e es]" + assume a0: "idx < length ?bs" + show "\i < length (?bs!idx). \k' pre red ps. pointer (?bs!idx!i) = PreRed (k', pre, red) ps \ red < i" + proof cases + assume a1: "idx=k" + consider (A) "item e \ set (items es)" | + (B) "item e \ set (items es) \ (\pre. pointer e = Null \ pointer e = Pre pre)" | + (C) "item e \ set (items es) \ \ (\pre. pointer e = Null \ pointer e = Pre pre)" + by blast + thus ?thesis + proof cases + case A + hence "bin_upd e es = es @ [e]" + using bin_upd_append by blast + thus ?thesis + using a1 assms(1-3,5) mono_red_ptr_def + by (metis length_append_singleton less_antisym nth_append nth_append_length nth_list_update_eq) + next + case B + hence "bin_upd e es = es" + using bin_upd_null_pre by blast + thus ?thesis + using a1 assms(1-3) mono_red_ptr_def by force + next + case C + then obtain i p ps where C: "i < length es" "item e = item (es!i)" "pointer e = PreRed p ps" + by (metis in_set_conv_nth items_def length_map nth_map pointer.exhaust) + show ?thesis + proof cases + assume "\p' ps'. pointer (es!i) = PreRed p' ps'" + hence "bin_upd e es = es" + using bin_upd_prered_nop C assms(4) by blast + thus ?thesis + using a1 assms(1-3) mono_red_ptr_def by (metis nth_list_update_eq) + next + assume "\ (\p' ps'. pointer (es!i) = PreRed p' ps')" + then obtain p' ps' where D: "pointer (es!i) = PreRed p' ps'" + by blast + have 0: "pointer (bin_upd e es!i) = PreRed p' (p#ps@ps') \ + (\j < length (bin_upd e es). i \ j \ bin_upd e es!j = es!j) \ + length (bin_upd e es) = length es" + using C D assms(4) bin_upd_prered_upd by blast + show ?thesis + proof (standard, standard, standard, standard, standard, standard, standard) + fix j k' pre red PS + assume a2: "j < length (?bs!idx)" + assume a3: "pointer (?bs!idx!j) = PreRed (k', pre, red) PS" + have 1: "?bs!idx = bin_upd e es" + by (simp add: a1 assms(2)) + show "red < j" + proof cases + assume a4: "i=j" + show ?thesis + using 0 1 C(1) D a3 a4 assms(1-3) unfolding mono_red_ptr_def by (metis pointer.inject(2)) + next + assume a4: "i\j" + thus ?thesis + using 0 1 a2 a3 assms(1) assms(2) assms(3) mono_red_ptr_def by force + qed + qed + qed + qed + next + assume a1: "idx\k" + show ?thesis + using a0 a1 assms(1) mono_red_ptr_def by fastforce + qed +qed + +lemma sound_mono_ptrs_bin_upds: + assumes "sound_ptrs \ bs" "mono_red_ptr bs" "k < length bs" "b = bs!k" "distinct (items b)" "distinct (items es)" + assumes "\e \ set es. sound_null_ptr e \ sound_pre_ptr \ bs k e \ sound_prered_ptr bs k e" + assumes "\e \ set es. \k' pre red ps. pointer e = PreRed (k', pre, red) ps \ red < length b" + shows "sound_ptrs \ (bs[k := bin_upds es b]) \ mono_red_ptr (bs[k := bin_upds es b])" + using assms +proof (induction es arbitrary: b bs) + case (Cons e es) + let ?bs = "bs[k := bin_upd e b]" + have 0: "sound_ptrs \ ?bs" + using sound_ptrs_bin_upd Cons.prems(1,3-5,7) by (metis list.set_intros(1)) + have 1: "mono_red_ptr ?bs" + using mono_red_ptr_bin_upd Cons.prems(2-5,8) by auto + have 2: "k < length ?bs" + using Cons.prems(3) by simp + have 3: "bin_upd e b = ?bs!k" + using Cons.prems(3) by simp + have 4: "\e' \ set es. sound_null_ptr e' \ sound_pre_ptr \ ?bs k e' \ sound_prered_ptr ?bs k e'" + using Cons.prems(3,4,7) length_bin_upd nth_item_bin_upd sound_pre_ptr_def sound_prered_ptr_def + by (smt (verit, ccfv_threshold) list.set_intros(2) nth_list_update order_less_le_trans) + have 5: "\e' \ set es. \k' pre red ps. pointer e' = PreRed (k', pre, red) ps \ red < length (bin_upd e b)" + by (meson Cons.prems(8) length_bin_upd order_less_le_trans set_subset_Cons subsetD) + have "sound_ptrs \ ((bs[k := bin_upd e b])[k := bin_upds es (bin_upd e b)]) \ + mono_red_ptr (bs[k := bin_upd e b, k := bin_upds es (bin_upd e b)])" + using Cons.IH[OF 0 1 2 3 _ _ 4 5] distinct_bin_upd Cons.prems(4,5,6) items_def by (metis distinct.simps(2) list.simps(9)) + thus ?case + by simp +qed simp + +lemma sound_mono_ptrs_Earley\<^sub>L_bin': + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "sound_ptrs \ bs" "\x \ bins bs. sound_item \ \ x" + assumes "mono_red_ptr bs" + assumes "nonempty_derives \" "wf_\ \" + shows "sound_ptrs \ (Earley\<^sub>L_bin' k \ \ bs i) \ mono_red_ptr (Earley\<^sub>L_bin' k \ \ bs i)" + using assms +proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) + case (Complete\<^sub>F k \ \ bs i x) + let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" + have x: "x \ set (items (bs ! k))" + using Complete\<^sub>F.hyps(1,2) by force + hence "\x \ set (items (Complete\<^sub>L k x bs i)). sound_item \ \ x" + using sound_Complete\<^sub>L Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items x + by (metis dual_order.refl) + hence sound: "\x \ bins ?bs'. sound_item \ \ x" + by (metis Complete\<^sub>F.prems(1,3) UnE bins_bins_upd wf_earley_input_elim) + have 0: "k < length bs" + using Complete\<^sub>F.prems(1) wf_earley_input_elim by auto + have 1: "\e \ set (Complete\<^sub>L k x bs i). sound_null_ptr e" + unfolding Complete\<^sub>L_def sound_null_ptr_def by auto + have 2: "\e \ set (Complete\<^sub>L k x bs i). sound_pre_ptr \ bs k e" + unfolding Complete\<^sub>L_def sound_pre_ptr_def by auto + { + fix e + assume a0: "e \ set (Complete\<^sub>L k x bs i)" + fix p ps k' pre red + assume a1: "pointer e = PreRed p ps" "(k', pre, red) \ set (p#ps)" + have "k' = item_origin x" + using a0 a1 unfolding Complete\<^sub>L_def by auto + moreover have "wf_item \ \ x" "item_end x = k" + using Complete\<^sub>F.prems(1) x wf_earley_input_elim wf_bins_kth_bin by blast+ + ultimately have 0: "k' \ k" + using wf_item_def by blast + have 1: "k' \ k" + proof (rule ccontr) + assume "\ k' \ k" + have "sound_item \ \ x" + using Complete\<^sub>F.prems(1,3) x kth_bin_sub_bins wf_earley_input_elim by (metis subset_eq) + moreover have "is_complete x" + using Complete\<^sub>F.hyps(3) by (auto simp: next_symbol_def split: if_splits) + moreover have "item_origin x = k" + using \\ k' \ k\ \k' = item_origin x\ by auto + ultimately show False + using impossible_complete_item Complete\<^sub>F.prems(1,5) wf_earley_input_elim \item_end x = k\ \wf_item \ \ x\ by blast + qed + have 2: "pre < length (bs!k')" + using a0 a1 index_filter_with_index_lt_length unfolding Complete\<^sub>L_def by (auto simp: items_def; fastforce) + have 3: "red < i+1" + using a0 a1 unfolding Complete\<^sub>L_def by auto + + have "item e = inc_item (item (bs!k'!pre)) k" + using a0 a1 0 2 Complete\<^sub>F.hyps(1,2,3) Complete\<^sub>F.prems(1) \k' = item_origin x\ unfolding Complete\<^sub>L_def + by (auto simp: items_def, metis filter_with_index_nth nth_map) + moreover have "is_complete (item (bs!k!red))" + using a0 a1 0 2 Complete\<^sub>F.hyps(1,2,3) Complete\<^sub>F.prems(1) \k' = item_origin x\ unfolding Complete\<^sub>L_def + by (auto simp: next_symbol_def items_def split: if_splits) + moreover have "item_origin (item (bs!k!red)) = item_end (item (bs!k'!pre))" + using a0 a1 0 2 Complete\<^sub>F.hyps(1,2,3) Complete\<^sub>F.prems(1) \k' = item_origin x\ unfolding Complete\<^sub>L_def + apply (clarsimp simp: items_def) + by (metis dual_order.strict_trans index_filter_with_index_lt_length items_def le_neq_implies_less nth_map nth_mem wf_bins_kth_bin wf_earley_input_elim) + moreover have "(\N. next_symbol (item (bs ! k' ! pre)) = Some N \ N = item_rule_head (item (bs ! k ! red)))" + using a0 a1 0 2 Complete\<^sub>F.hyps(1,2,3) Complete\<^sub>F.prems(1) \k' = item_origin x\ unfolding Complete\<^sub>L_def + by (auto simp: items_def, metis (mono_tags, lifting) filter_with_index_P filter_with_index_nth nth_map) + ultimately have 4: "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" + unfolding completes_def by blast + have "k' < k" "pre < length (bs!k')" "red < i+1" "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" + using 0 1 2 3 4 by simp_all + } + hence "\e \ set (Complete\<^sub>L k x bs i). \p ps k' pre red. pointer e = PreRed p ps \ (k', pre, red) \ set (p#ps) \ + k' < k \ pre < length (bs!k') \ red < i+1 \ completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" + by force + hence 3: "\e \ set (Complete\<^sub>L k x bs i). sound_prered_ptr bs k e" + unfolding sound_prered_ptr_def using Complete\<^sub>F.hyps(1) items_def by (smt (verit) discrete dual_order.strict_trans1 leI length_map) + have 4: "distinct (items (Complete\<^sub>L k x bs i))" + using distinct_Complete\<^sub>L x Complete\<^sub>F.prems(1) wf_earley_input_elim wf_bin_def wf_bin_items_def wf_bins_def wf_item_def + by (metis order_le_less_trans) + have "sound_ptrs \ ?bs' \ mono_red_ptr ?bs'" + using sound_mono_ptrs_bin_upds[OF Complete\<^sub>F.prems(2) Complete\<^sub>F.prems(4) 0] 1 2 3 4 sound_prered_ptr_def + Complete\<^sub>F.prems(1) bins_upd_def wf_earley_input_elim wf_bin_def wf_bins_def + by (smt (verit, ccfv_SIG) list.set_intros(1)) + moreover have "(k, \, \, ?bs') \ wf_earley_input" + using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast + ultimately have "sound_ptrs \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) \ mono_red_ptr (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" + using Complete\<^sub>F.IH Complete\<^sub>F.prems(4-6) sound by blast + thus ?case + using Complete\<^sub>F.hyps by simp +next + case (Scan\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" + have "x \ set (items (bs ! k))" + using Scan\<^sub>F.hyps(1,2) by force + hence "\x \ set (items (Scan\<^sub>L k \ a x i)). sound_item \ \ x" + using sound_Scan\<^sub>L Scan\<^sub>F.hyps(3,5) Scan\<^sub>F.prems(1,2,3) wf_earley_input_elim wf_bins_impl_wf_items wf_bins_impl_wf_items by fast + hence sound: "\x \ bins ?bs'. sound_item \ \ x" + using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1,3) bins_bins_upd wf_earley_input_elim + by (metis UnE add_less_cancel_right) + have 0: "k+1 < length bs" + using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1) wf_earley_input_elim by force + have 1: "\e \ set (Scan\<^sub>L k \ a x i). sound_null_ptr e" + unfolding Scan\<^sub>L_def sound_null_ptr_def by auto + have 2: "\e \ set (Scan\<^sub>L k \ a x i). sound_pre_ptr \ bs (k+1) e" + using Scan\<^sub>F.hyps(1,2,3) unfolding sound_pre_ptr_def Scan\<^sub>L_def scans_def items_def by auto + have 3: "\e \ set (Scan\<^sub>L k \ a x i). sound_prered_ptr bs (k+1) e" + unfolding Scan\<^sub>L_def sound_prered_ptr_def by simp + have 4: "distinct (items (Scan\<^sub>L k \ a x i))" + using distinct_Scan\<^sub>L by fast + have "sound_ptrs \ ?bs' \ mono_red_ptr ?bs'" + using sound_mono_ptrs_bin_upds[OF Scan\<^sub>F.prems(2) Scan\<^sub>F.prems(4) 0] 0 1 2 3 4 sound_prered_ptr_def + Scan\<^sub>F.prems(1) bins_upd_def wf_earley_input_elim wf_bin_def wf_bins_def + by (smt (verit, ccfv_threshold) list.set_intros(1)) + moreover have "(k, \, \, ?bs') \ wf_earley_input" + using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis + ultimately have "sound_ptrs \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) \ mono_red_ptr (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" + using Scan\<^sub>F.IH Scan\<^sub>F.prems(4-6) sound by blast + thus ?case + using Scan\<^sub>F.hyps by simp +next + case (Predict\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" + have "x \ set (items (bs ! k))" + using Predict\<^sub>F.hyps(1,2) by force + hence "\x \ set (items(Predict\<^sub>L k \ a)). sound_item \ \ x" + using sound_Predict\<^sub>L Predict\<^sub>F.hyps(3) Predict\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items by fast + hence sound: "\x \ bins ?bs'. sound_item \ \ x" + using Predict\<^sub>F.prems(1,3) UnE bins_bins_upd wf_earley_input_elim by metis + have 0: "k < length bs" + using Predict\<^sub>F.prems(1) wf_earley_input_elim by force + have 1: "\e \ set (Predict\<^sub>L k \ a). sound_null_ptr e" + unfolding sound_null_ptr_def Predict\<^sub>L_def predicts_def by (auto simp: init_item_def) + have 2: "\e \ set (Predict\<^sub>L k \ a). sound_pre_ptr \ bs k e" + unfolding sound_pre_ptr_def Predict\<^sub>L_def by simp + have 3: "\e \ set (Predict\<^sub>L k \ a). sound_prered_ptr bs k e" + unfolding sound_prered_ptr_def Predict\<^sub>L_def by simp + have 4: "distinct (items (Predict\<^sub>L k \ a))" + using Predict\<^sub>F.prems(6) distinct_Predict\<^sub>L by fast + have "sound_ptrs \ ?bs' \ mono_red_ptr ?bs'" + using sound_mono_ptrs_bin_upds[OF Predict\<^sub>F.prems(2) Predict\<^sub>F.prems(4) 0] 0 1 2 3 4 sound_prered_ptr_def + Predict\<^sub>F.prems(1) bins_upd_def wf_earley_input_elim wf_bin_def wf_bins_def + by (smt (verit, ccfv_threshold) list.set_intros(1)) + moreover have "(k, \, \, ?bs') \ wf_earley_input" + using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis + ultimately have "sound_ptrs \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) \ mono_red_ptr (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" + using Predict\<^sub>F.IH Predict\<^sub>F.prems(4-6) sound by blast + thus ?case + using Predict\<^sub>F.hyps by simp +qed simp_all + +lemma sound_mono_ptrs_Earley\<^sub>L_bin: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "sound_ptrs \ bs" "\x \ bins bs. sound_item \ \ x" + assumes "mono_red_ptr bs" + assumes "nonempty_derives \" "wf_\ \" + shows "sound_ptrs \ (Earley\<^sub>L_bin k \ \ bs) \ mono_red_ptr (Earley\<^sub>L_bin k \ \ bs)" + using assms sound_mono_ptrs_Earley\<^sub>L_bin' Earley\<^sub>L_bin_def by metis + +lemma sound_ptrs_Init\<^sub>L: + "sound_ptrs \ (Init\<^sub>L \ \)" + unfolding sound_ptrs_def sound_null_ptr_def sound_pre_ptr_def sound_prered_ptr_def + predicts_def scans_def completes_def Init\<^sub>L_def + by (auto simp: init_item_def less_Suc_eq_0_disj) + +lemma mono_red_ptr_Init\<^sub>L: + "mono_red_ptr (Init\<^sub>L \ \)" + unfolding mono_red_ptr_def Init\<^sub>L_def + by (auto simp: init_item_def less_Suc_eq_0_disj) + +lemma sound_mono_ptrs_Earley\<^sub>L_bins: + assumes "k \ length \" "wf_\ \" "nonempty_derives \" "wf_\ \" + shows "sound_ptrs \ (Earley\<^sub>L_bins k \ \) \ mono_red_ptr (Earley\<^sub>L_bins k \ \)" + using assms +proof (induction k) + case 0 + have "(0, \, \, (Init\<^sub>L \ \)) \ wf_earley_input" + using assms(2) wf_earley_input_Init\<^sub>L by blast + moreover have "\x \ bins (Init\<^sub>L \ \). sound_item \ \ x" + by (metis Init\<^sub>L_eq_Init\<^sub>F Init\<^sub>F_sub_Earley sound_Earley subsetD wf_Earley) + ultimately show ?case + using sound_mono_ptrs_Earley\<^sub>L_bin sound_ptrs_Init\<^sub>L mono_red_ptr_Init\<^sub>L "0.prems"(2,3) by fastforce +next + case (Suc k) + have "(Suc k, \, \, Earley\<^sub>L_bins k \ \) \ wf_earley_input" + by (simp add: Suc.prems(1) Suc_leD assms(2) wf_earley_input_intro) + moreover have "sound_ptrs \ (Earley\<^sub>L_bins k \ \)" + using Suc by simp + moreover have "\x \ bins (Earley\<^sub>L_bins k \ \). sound_item \ \ x" + by (meson Suc.prems(1) Suc_leD Earley\<^sub>L_bins_sub_Earley\<^sub>F_bins Earley\<^sub>F_bins_sub_Earley assms(2) + sound_Earley subsetD wf_bins_Earley\<^sub>L_bins wf_bins_impl_wf_items) + ultimately show ?case + using Suc.prems(1,3,4) sound_mono_ptrs_Earley\<^sub>L_bin Suc.IH by fastforce +qed + +lemma sound_mono_ptrs_Earley\<^sub>L: + assumes "wf_\ \" "nonempty_derives \" + shows "sound_ptrs \ (Earley\<^sub>L \ \) \ mono_red_ptr (Earley\<^sub>L \ \)" + using assms sound_mono_ptrs_Earley\<^sub>L_bins Earley\<^sub>L_def by (metis dual_order.refl) + + +subsection \Common Definitions\ + +datatype 'a tree = + Leaf 'a + | Branch 'a "'a tree list" + +fun yield_tree :: "'a tree \ 'a sentence" where + "yield_tree (Leaf a) = [a]" +| "yield_tree (Branch _ ts) = concat (map yield_tree ts)" + +fun root_tree :: "'a tree \ 'a" where + "root_tree (Leaf a) = a" +| "root_tree (Branch N _) = N" + +fun wf_rule_tree :: "'a cfg \ 'a tree \ bool" where + "wf_rule_tree _ (Leaf a) \ True" +| "wf_rule_tree \ (Branch N ts) \ ( + (\r \ set (\ \). N = rule_head r \ map root_tree ts = rule_body r) \ + (\t \ set ts. wf_rule_tree \ t))" + +fun wf_item_tree :: "'a cfg \ 'a item \ 'a tree \ bool" where + "wf_item_tree \ _ (Leaf a) \ True" +| "wf_item_tree \ x (Branch N ts) \ ( + N = item_rule_head x \ map root_tree ts = take (item_dot x) (item_rule_body x) \ + (\t \ set ts. wf_rule_tree \ t))" + +definition wf_yield_tree :: "'a sentence \ 'a item \ 'a tree \ bool" where + "wf_yield_tree \ x t \ yield_tree t = slice (item_origin x) (item_end x) \" + +datatype 'a forest = + FLeaf 'a + | FBranch 'a "'a forest list list" + +fun combinations :: "'a list list \ 'a list list" where + "combinations [] = [[]]" +| "combinations (xs#xss) = [ x#cs . x <- xs, cs <- combinations xss ]" + +fun trees :: "'a forest \ 'a tree list" where + "trees (FLeaf a) = [Leaf a]" +| "trees (FBranch N fss) = ( + let tss = (map (\fs. concat (map (\f. trees f) fs)) fss) in + map (\ts. Branch N ts) (combinations tss) + )" + +lemma list_comp_flatten: + "[ f xs . xs <- [ g xs ys . xs <- as, ys <- bs ] ] = [ f (g xs ys) . xs <- as, ys <- bs ]" + by (induction as) auto + +lemma list_comp_flatten_Cons: + "[ x#xs . x <- as, xs <- [ xs @ ys. xs <- bs, ys <- cs ] ] = [ x#xs@ys. x <- as, xs <- bs, ys <- cs ]" + by (induction as) (auto simp: list_comp_flatten) + +lemma list_comp_flatten_append: + "[ xs@ys . xs <- [ x#xs . x <- as, xs <- bs ], ys <- cs ] = [ x#xs@ys . x <- as, xs <- bs, ys <- cs ]" + by (induction as) (auto simp: o_def, meson append_Cons map_eq_conv) + +lemma combinations_append: + "combinations (xss @ yss) = [ xs @ ys . xs <- combinations xss, ys <- combinations yss ]" + by (induction xss) (auto simp: list_comp_flatten_Cons list_comp_flatten_append map_idI) + +lemma trees_append: + "trees (FBranch N (xss @ yss)) = ( + let xtss = (map (\xs. concat (map (\f. trees f) xs)) xss) in + let ytss = (map (\ys. concat (map (\f. trees f) ys)) yss) in + map (\ts. Branch N ts) [ xs @ ys . xs <- combinations xtss, ys <- combinations ytss ])" + using combinations_append by (metis map_append trees.simps(2)) + +lemma trees_append_singleton: + "trees (FBranch N (xss @ [ys])) = ( + let xtss = (map (\xs. concat (map (\f. trees f) xs)) xss) in + let ytss = [concat (map trees ys)] in + map (\ts. Branch N ts) [ xs @ ys . xs <- combinations xtss, ys <- combinations ytss ])" + by (subst trees_append, simp) + +lemma trees_append_single_singleton: + "trees (FBranch N (xss @ [[y]])) = ( + let xtss = (map (\xs. concat (map (\f. trees f) xs)) xss) in + map (\ts. Branch N ts) [ xs @ ys . xs <- combinations xtss, ys <- [ [t] . t <- trees y ] ])" + by (subst trees_append_singleton, auto) + + +subsection \foldl lemmas\ + +lemma foldl_add_nth: + "k < length xs \ foldl (+) z (map length (take k xs)) + length (xs!k) = foldl (+) z (map length (take (k+1) xs))" +proof (induction xs arbitrary: k z) + case (Cons x xs) + then show ?case + proof (cases "k = 0") + case False + thus ?thesis + using Cons by (auto simp add: take_Cons') + qed simp +qed simp + +lemma foldl_acc_mono: + "a \ b \ foldl (+) a xs \ foldl (+) b xs" for a :: nat + by (induction xs arbitrary: a b) auto + +lemma foldl_ge_z_nth: + "j < length xs \ z + length (xs!j) \ foldl (+) z (map length (take (j+1) xs))" +proof (induction xs arbitrary: j z) + case (Cons x xs) + show ?case + proof (cases "j = 0") + case False + have "z + length ((x # xs) ! j) = z + length (xs!(j-1))" + using False by simp + also have "... \ foldl (+) z (map length (take (j-1+1) xs))" + using Cons False by (metis add_diff_inverse_nat length_Cons less_one nat_add_left_cancel_less plus_1_eq_Suc) + also have "... = foldl (+) z (map length (take j xs))" + using False by simp + also have "... \ foldl (+) (z + length x) (map length (take j xs))" + using foldl_acc_mono by force + also have "... = foldl (+) z (map length (take (j+1) (x#xs)))" + by simp + finally show ?thesis + by blast + qed simp +qed simp + +lemma foldl_add_nth_ge: + "i \ j \ j < length xs \ foldl (+) z (map length (take i xs)) + length (xs!j) \ foldl (+) z (map length (take (j+1) xs))" +proof (induction xs arbitrary: i j z) + case (Cons x xs) + show ?case + proof (cases "i = 0") + case True + have "foldl (+) z (map length (take i (x # xs))) + length ((x # xs) ! j) = z + length ((x # xs) ! j)" + using True by simp + also have "... \ foldl (+) z (map length (take (j+1) (x#xs)))" + using foldl_ge_z_nth Cons.prems(2) by blast + finally show ?thesis + by blast + next + case False + have "i-1 \ j-1" + by (simp add: Cons.prems(1) diff_le_mono) + have "j-1 < length xs" + using Cons.prems(1,2) False by fastforce + have "foldl (+) z (map length (take i (x # xs))) + length ((x # xs) ! j) = + foldl (+) (z + length x) (map length (take (i-1) xs)) + length ((x#xs)!j)" + using False by (simp add: take_Cons') + also have "... = foldl (+) (z + length x) (map length (take (i-1) xs)) + length (xs!(j-1))" + using Cons.prems(1) False by auto + also have "... \ foldl (+) (z + length x) (map length (take (j-1+1) xs))" + using Cons.IH \i - 1 \ j - 1\ \j - 1 < length xs\ by blast + also have "... = foldl (+) (z + length x) (map length (take j xs))" + using Cons.prems(1) False by fastforce + also have "... = foldl (+) z (map length (take (j+1) (x#xs)))" + by fastforce + finally show ?thesis + by blast + qed +qed simp + +lemma foldl_ge_acc: + "foldl (+) z (map length xs) \ z" + by (induction xs arbitrary: z) (auto elim: add_leE) + +lemma foldl_take_mono: + "i \ j \ foldl (+) z (map length (take i xs)) \ foldl (+) z (map length (take j xs))" +proof (induction xs arbitrary: z i j) + case (Cons x xs) + show ?case + proof (cases "i = 0") + case True + have "foldl (+) z (map length (take i (x # xs))) = z" + using True by simp + also have "... \ foldl (+) z (map length (take j (x # xs)))" + by (simp add: foldl_ge_acc) + ultimately show ?thesis + by simp + next + case False + then show ?thesis + using Cons by (simp add: take_Cons') + qed +qed simp + + +subsection \Parse tree\ + +partial_function (option) build_tree' :: "'a bins \ 'a sentence \ nat \ nat \ 'a tree option" where + "build_tree' bs \ k i = ( + let e = bs!k!i in ( + case pointer e of + Null \ Some (Branch (item_rule_head (item e)) []) \\start building sub-tree\ + | Pre pre \ ( \\add sub-tree starting from terminal\ + do { + t \ build_tree' bs \ (k-1) pre; + case t of + Branch N ts \ Some (Branch N (ts @ [Leaf (\!(k-1))])) + | _ \ undefined \\impossible case\ + }) + | PreRed (k', pre, red) _ \ ( \\add sub-tree starting from non-terminal\ + do { + t \ build_tree' bs \ k' pre; + case t of + Branch N ts \ + do { + t \ build_tree' bs \ k red; + Some (Branch N (ts @ [t])) + } + | _ \ undefined \\impossible case\ + }) + ))" + +declare build_tree'.simps [code] + +definition build_tree :: "'a cfg \ 'a sentence \ 'a bins \ 'a tree option" where + "build_tree \ \ bs = ( + let k = length bs - 1 in ( + case filter_with_index (\x. is_finished \ \ x) (items (bs!k)) of + [] \ None + | (_, i)#_ \ build_tree' bs \ k i + ))" + +lemma build_tree'_simps[simp]: + "e = bs!k!i \ pointer e = Null \ build_tree' bs \ k i = Some (Branch (item_rule_head (item e)) [])" + "e = bs!k!i \ pointer e = Pre pre \ build_tree' bs \ (k-1) pre = None \ + build_tree' bs \ k i = None" + "e = bs!k!i \ pointer e = Pre pre \ build_tree' bs \ (k-1) pre = Some (Branch N ts) \ + build_tree' bs \ k i = Some (Branch N (ts @ [Leaf (\!(k-1))]))" + "e = bs!k!i \ pointer e = Pre pre \ build_tree' bs \ (k-1) pre = Some (Leaf a) \ + build_tree' bs \ k i = undefined" + "e = bs!k!i \ pointer e = PreRed (k', pre, red) reds \ build_tree' bs \ k' pre = None \ + build_tree' bs \ k i = None" + "e = bs!k!i \ pointer e = PreRed (k', pre, red) reds \ build_tree' bs \ k' pre = Some (Branch N ts) \ + build_tree' bs \ k red = None \ build_tree' bs \ k i = None" + "e = bs!k!i \ pointer e = PreRed (k', pre, red) reds \ build_tree' bs \ k' pre = Some (Leaf a) \ + build_tree' bs \ k i = undefined" + "e = bs!k!i \ pointer e = PreRed (k', pre, red) reds \ build_tree' bs \ k' pre = Some (Branch N ts) \ + build_tree' bs \ k red = Some t \ + build_tree' bs \ k i = Some (Branch N (ts @ [t]))" + by (subst build_tree'.simps, simp)+ + +definition wf_tree_input :: "('a bins \ 'a sentence \ nat \ nat) set" where + "wf_tree_input = { + (bs, \, k, i) | bs \ k i. + sound_ptrs \ bs \ + mono_red_ptr bs \ + k < length bs \ + i < length (bs!k) + }" + +fun build_tree'_measure :: "('a bins \ 'a sentence \ nat \ nat) \ nat" where + "build_tree'_measure (bs, \, k, i) = foldl (+) 0 (map length (take k bs)) + i" + +lemma wf_tree_input_pre: + assumes "(bs, \, k, i) \ wf_tree_input" + assumes "e = bs!k!i" "pointer e = Pre pre" + shows "(bs, \, (k-1), pre) \ wf_tree_input" + using assms unfolding wf_tree_input_def + using less_imp_diff_less nth_mem by (fastforce simp: sound_ptrs_def sound_pre_ptr_def) + +lemma wf_tree_input_prered_pre: + assumes "(bs, \, k, i) \ wf_tree_input" + assumes "e = bs!k!i" "pointer e = PreRed (k', pre, red) ps" + shows "(bs, \, k', pre) \ wf_tree_input" + using assms unfolding wf_tree_input_def + apply (auto simp: sound_ptrs_def sound_prered_ptr_def) + apply metis+ + apply (metis dual_order.strict_trans nth_mem) + by (metis nth_mem) + +lemma wf_tree_input_prered_red: + assumes "(bs, \, k, i) \ wf_tree_input" + assumes "e = bs!k!i" "pointer e = PreRed (k', pre, red) ps" + shows "(bs, \, k, red) \ wf_tree_input" + using assms unfolding wf_tree_input_def + apply (auto simp add: sound_ptrs_def sound_prered_ptr_def) + apply (metis nth_mem)+ + done + +lemma build_tree'_induct: + assumes "(bs, \, k, i) \ wf_tree_input" + assumes "\bs \ k i. + (\e pre. e = bs!k!i \ pointer e = Pre pre \ P bs \ (k-1) pre) \ + (\e k' pre red ps. e = bs!k!i \ pointer e = PreRed (k', pre, red) ps \ P bs \ k' pre) \ + (\e k' pre red ps. e = bs!k!i \ pointer e = PreRed (k', pre, red) ps \ P bs \ k red) \ + P bs \ k i" + shows "P bs \ k i" + using assms(1) +proof (induction n\"build_tree'_measure (bs, \, k, i)" arbitrary: k i rule: nat_less_induct) + case 1 + obtain e where entry: "e = bs!k!i" + by simp + consider (Null) "pointer e = Null" + | (Pre) "\pre. pointer e = Pre pre" + | (PreRed) "\k' pre red reds. pointer e = PreRed (k', pre, red) reds" + by (metis pointer.exhaust surj_pair) + thus ?case + proof cases + case Null + thus ?thesis + using assms(2) entry by fastforce + next + case Pre + then obtain pre where pre: "pointer e = Pre pre" + by blast + define n where n: "n = build_tree'_measure (bs, \, (k-1), pre)" + have "0 < k" "pre < length (bs!(k-1))" + using 1(2) entry pre unfolding wf_tree_input_def sound_ptrs_def sound_pre_ptr_def + by (smt (verit) mem_Collect_eq nth_mem prod.inject)+ + have "k < length bs" + using 1(2) unfolding wf_tree_input_def by blast+ + have "foldl (+) 0 (map length (take k bs)) + i - (foldl (+) 0 (map length (take (k-1) bs)) + pre) = + foldl (+) 0 (map length (take (k-1) bs)) + length (bs!(k-1)) + i - (foldl (+) 0 (map length (take (k-1) bs)) + pre)" + using foldl_add_nth[of \k-1\ bs 0] by (simp add: \0 < k\ \k < length bs\ less_imp_diff_less) + also have "... = length (bs!(k-1)) + i - pre" + by simp + also have "... > 0" + using \pre < length (bs!(k-1))\ by auto + finally have "build_tree'_measure (bs, \, k, i) - build_tree'_measure (bs, \, (k-1), pre) > 0" + by simp + hence "P bs \ (k-1) pre" + using 1 n wf_tree_input_pre entry pre zero_less_diff by blast + thus ?thesis + using assms(2) entry pre pointer.distinct(5) pointer.inject(1) by presburger + next + case PreRed + then obtain k' pre red ps where prered: "pointer e = PreRed (k', pre, red) ps" + by blast + have "k' < k" "pre < length (bs!k')" + using 1(2) entry prered unfolding wf_tree_input_def sound_ptrs_def sound_prered_ptr_def + apply simp_all + apply (metis nth_mem)+ + done + have "red < i" + using 1(2) entry prered unfolding wf_tree_input_def mono_red_ptr_def by blast + have "k < length bs" "i < length (bs!k)" + using 1(2) unfolding wf_tree_input_def by blast+ + define n_pre where n_pre: "n_pre = build_tree'_measure (bs, \, k', pre)" + have "0 < length (bs!k') + i - pre" + by (simp add: \pre < length (bs!k')\ add.commute trans_less_add2) + also have "... = foldl (+) 0 (map length (take k' bs)) + length (bs!k') + i - (foldl (+) 0 (map length (take k' bs)) + pre)" + by simp + also have "... \ foldl (+) 0 (map length (take (k'+1) bs)) + i - (foldl (+) 0 (map length (take k' bs)) + pre)" + using foldl_add_nth_ge[of k' k' bs 0] \k < length bs\ \k' < k\ by simp + also have "... \ foldl (+) 0 (map length (take k bs)) + i - (foldl (+) 0 (map length (take k' bs)) + pre)" + using foldl_take_mono by (metis Suc_eq_plus1 Suc_leI \k' < k\ add.commute add_le_cancel_left diff_le_mono) + finally have "build_tree'_measure (bs, \, k, i) - build_tree'_measure (bs, \, k', pre) > 0" + by simp + hence x: "P bs \ k' pre" + using 1(1) zero_less_diff by (metis "1.prems" entry prered wf_tree_input_prered_pre) + define n_red where n_red: "n_red = build_tree'_measure (bs, \, k, red)" + have "build_tree'_measure (bs, \, k, i) - build_tree'_measure (bs, \, k, red) > 0" + using \red < i\ by simp + hence y: "P bs \ k red" + using "1.hyps" "1.prems" entry prered wf_tree_input_prered_red zero_less_diff by blast + show ?thesis + using assms(2) x y entry prered + by (smt (verit, best) Pair_inject filter_cong pointer.distinct(5) pointer.inject(2)) + qed +qed + +lemma build_tree'_termination: + assumes "(bs, \, k, i) \ wf_tree_input" + shows "\N ts. build_tree' bs \ k i = Some (Branch N ts)" +proof - + have "\N ts. build_tree' bs \ k i = Some (Branch N ts)" + apply (induction rule: build_tree'_induct[OF assms(1)]) + subgoal premises IH for bs \ k i + proof - + define e where entry: "e = bs!k!i" + consider (Null) "pointer e = Null" + | (Pre) "\pre. pointer e = Pre pre" + | (PreRed) "\k' pre red ps. pointer e = PreRed (k', pre, red) ps" + by (metis pointer.exhaust surj_pair) + thus ?thesis + proof cases + case Null + thus ?thesis + using build_tree'_simps(1) entry by simp + next + case Pre + then obtain pre where pre: "pointer e = Pre pre" + by blast + obtain N ts where Nts: "build_tree' bs \ (k-1) pre = Some (Branch N ts)" + using IH(1) entry pre by blast + have "build_tree' bs \ k i = Some (Branch N (ts @ [Leaf (\!(k-1))]))" + using build_tree'_simps(3) entry pre Nts by simp + thus ?thesis + by simp + next + case PreRed + then obtain k' pre red ps where prered: "pointer e = PreRed (k', pre, red) ps" + by blast + then obtain N ts where Nts: "build_tree' bs \ k' pre = Some (Branch N ts)" + using IH(2) entry prered by blast + obtain t_red where t_red: "build_tree' bs \ k red = Some t_red" + using IH(3) entry prered Nts by (metis option.exhaust) + have "build_tree' bs \ k i = Some (Branch N (ts @ [t_red]))" + using build_tree'_simps(8) entry prered Nts t_red by auto + thus ?thesis + by blast + qed + qed + done + thus ?thesis + by blast +qed + +lemma wf_item_tree_build_tree': + assumes "(bs, \, k, i) \ wf_tree_input" + assumes "wf_bins \ \ bs" + assumes "k < length bs" "i < length (bs!k)" + assumes "build_tree' bs \ k i = Some t" + shows "wf_item_tree \ (item (bs!k!i)) t" +proof - + have "wf_item_tree \ (item (bs!k!i)) t" + using assms + apply (induction arbitrary: t rule: build_tree'_induct[OF assms(1)]) + subgoal premises prems for bs \ k i t + proof - + define e where entry: "e = bs!k!i" + consider (Null) "pointer e = Null" + | (Pre) "\pre. pointer e = Pre pre" + | (PreRed) "\k' pre red ps. pointer e = PreRed (k', pre, red) ps" + by (metis pointer.exhaust surj_pair) + thus ?thesis + proof cases + case Null + hence "build_tree' bs \ k i = Some (Branch (item_rule_head (item e)) [])" + using entry by simp + have simp: "t = Branch (item_rule_head (item e)) []" + using build_tree'_simps(1) Null prems(8) entry by simp + have "sound_ptrs \ bs" + using prems(4) unfolding wf_tree_input_def by blast + hence "predicts (item e)" + using Null prems(6,7) nth_mem entry unfolding sound_ptrs_def sound_null_ptr_def by blast + hence "item_dot (item e) = 0" + unfolding predicts_def by blast + thus ?thesis + using simp entry by simp + next + case Pre + then obtain pre where pre: "pointer e = Pre pre" + by blast + obtain N ts where Nts: "build_tree' bs \ (k-1) pre = Some (Branch N ts)" + using build_tree'_termination entry pre prems(4) wf_tree_input_pre by blast + have simp: "build_tree' bs \ k i = Some (Branch N (ts @ [Leaf (\!(k-1))]))" + using build_tree'_simps(3) entry pre Nts by simp + have "sound_ptrs \ bs" + using prems(4) unfolding wf_tree_input_def by blast + hence "pre < length (bs!(k-1))" + using entry pre prems(6,7) unfolding sound_ptrs_def sound_pre_ptr_def by (metis nth_mem) + moreover have "k-1 < length bs" + by (simp add: prems(6) less_imp_diff_less) + ultimately have IH: "wf_item_tree \ (item (bs!(k-1)!pre)) (Branch N ts)" + using prems(1,2,4,5) entry pre Nts by (meson wf_tree_input_pre) + have scans: "scans \ k (item (bs!(k-1)!pre)) (item e)" + using entry pre prems(6-7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_pre_ptr_def by simp + hence *: + "item_rule_head (item (bs!(k-1)!pre)) = item_rule_head (item e)" + "item_rule_body (item (bs!(k-1)!pre)) = item_rule_body (item e)" + "item_dot (item (bs!(k-1)!pre)) + 1 = item_dot (item e)" + "next_symbol (item (bs!(k-1)!pre)) = Some (\!(k-1))" + unfolding scans_def inc_item_def by (simp_all add: item_rule_head_def item_rule_body_def) + have "map root_tree (ts @ [Leaf (\!(k-1))]) = map root_tree ts @ [\!(k-1)]" + by simp + also have "... = take (item_dot (item (bs!(k-1)!pre))) (item_rule_body (item (bs!(k-1)!pre))) @ [\!(k-1)]" + using IH by simp + also have "... = take (item_dot (item (bs!(k-1)!pre))) (item_rule_body (item e)) @ [\!(k-1)]" + using *(2) by simp + also have "... = take (item_dot (item e)) (item_rule_body (item e))" + using *(2-4) by (auto simp: next_symbol_def is_complete_def split: if_splits; metis leI take_Suc_conv_app_nth) + finally have "map root_tree (ts @ [Leaf (\!(k-1))]) = take (item_dot (item e)) (item_rule_body (item e))" . + hence "wf_item_tree \ (item e) (Branch N (ts @ [Leaf (\!(k-1))]))" + using IH *(1) by simp + thus ?thesis + using entry prems(8) simp by auto + next + case PreRed + then obtain k' pre red ps where prered: "pointer e = PreRed (k', pre, red) ps" + by blast + obtain N ts where Nts: "build_tree' bs \ k' pre = Some (Branch N ts)" + using build_tree'_termination entry prems(4) prered wf_tree_input_prered_pre by blast + obtain N_red ts_red where Nts_red: "build_tree' bs \ k red = Some (Branch N_red ts_red)" + using build_tree'_termination entry prems(4) prered wf_tree_input_prered_red by blast + have simp: "build_tree' bs \ k i = Some (Branch N (ts @ [Branch N_red ts_red]))" + using build_tree'_simps(8) entry prered Nts Nts_red by auto + have "sound_ptrs \ bs" + using prems(4) wf_tree_input_def by fastforce + have bounds: "k' < k" "pre < length (bs!k')" "red < length (bs!k)" + using prered entry prems(6,7) \sound_ptrs \ bs\ + unfolding sound_prered_ptr_def sound_ptrs_def by fastforce+ + have completes: "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" + using prered entry prems(6,7) \sound_ptrs \ bs\ + unfolding sound_ptrs_def sound_prered_ptr_def by fastforce + have *: + "item_rule_head (item (bs!k'!pre)) = item_rule_head (item e)" + "item_rule_body (item (bs!k'!pre)) = item_rule_body (item e)" + "item_dot (item (bs!k'!pre)) + 1 = item_dot (item e)" + "next_symbol (item (bs!k'!pre)) = Some (item_rule_head (item (bs!k!red)))" + "is_complete (item (bs!k!red))" + using completes unfolding completes_def inc_item_def + by (auto simp: item_rule_head_def item_rule_body_def is_complete_def) + have "(bs, \, k', pre) \ wf_tree_input" + using wf_tree_input_prered_pre[OF prems(4) entry prered] by blast + hence IH_pre: "wf_item_tree \ (item (bs!k'!pre)) (Branch N ts)" + using prems(2)[OF entry prered _ prems(5)] Nts bounds(1,2) order_less_trans prems(6) by blast + have "(bs, \, k, red) \ wf_tree_input" + using wf_tree_input_prered_red[OF prems(4) entry prered] by blast + hence IH_r: "wf_item_tree \ (item (bs!k!red)) (Branch N_red ts_red)" + using bounds(3) entry prems(3,5,6) prered Nts_red by blast + have "map root_tree (ts @ [Branch N_red ts_red]) = map root_tree ts @ [root_tree (Branch N_red ts)]" + by simp + also have "... = take (item_dot (item (bs!k'!pre))) (item_rule_body (item (bs!k'!pre))) @ [root_tree (Branch N_red ts)]" + using IH_pre by simp + also have "... = take (item_dot (item (bs!k'!pre))) (item_rule_body (item (bs!k'!pre))) @ [item_rule_head (item (bs!k!red))]" + using IH_r by simp + also have "... = take (item_dot (item e)) (item_rule_body (item e))" + using * by (auto simp: next_symbol_def is_complete_def split: if_splits; metis leI take_Suc_conv_app_nth) + finally have roots: "map root_tree (ts @ [Branch N_red ts]) = take (item_dot (item e)) (item_rule_body (item e))" by simp + have "wf_item \ \ (item (bs!k!red))" + using prems(5,6) bounds(3) unfolding wf_bins_def wf_bin_def wf_bin_items_def by (auto simp: items_def) + moreover have "N_red = item_rule_head (item (bs!k!red))" + using IH_r by fastforce + moreover have "map root_tree ts_red = item_rule_body (item (bs!k!red))" + using IH_r *(5) by (auto simp: is_complete_def) + ultimately have "\r \ set (\ \). N_red = rule_head r \ map root_tree ts_red = rule_body r" + unfolding wf_item_def item_rule_body_def item_rule_head_def by blast + hence "wf_rule_tree \ (Branch N_red ts_red)" + using IH_r by simp + hence "wf_item_tree \ (item (bs!k!i)) (Branch N (ts @ [Branch N_red ts_red]))" + using "*"(1) roots IH_pre entry by simp + thus ?thesis + using Nts_red prems(8) simp by auto + qed + qed + done + thus ?thesis + using assms(2) by blast +qed + +lemma wf_yield_tree_build_tree': + assumes "(bs, \, k, i) \ wf_tree_input" + assumes "wf_bins \ \ bs" + assumes "k < length bs" "i < length (bs!k)" "k \ length \" + assumes "build_tree' bs \ k i = Some t" + shows "wf_yield_tree \ (item (bs!k!i)) t" +proof - + have "wf_yield_tree \ (item (bs!k!i)) t" + using assms + apply (induction arbitrary: t rule: build_tree'_induct[OF assms(1)]) + subgoal premises prems for bs \ k i t + proof - + define e where entry: "e = bs!k!i" + consider (Null) "pointer e = Null" + | (Pre) "\pre. pointer e = Pre pre" + | (PreRed) "\k' pre red reds. pointer e = PreRed (k', pre, red) reds" + by (metis pointer.exhaust surj_pair) + thus ?thesis + proof cases + case Null + hence "build_tree' bs \ k i = Some (Branch (item_rule_head (item e)) [])" + using entry by simp + have simp: "t = Branch (item_rule_head (item e)) []" + using build_tree'_simps(1) Null prems(9) entry by simp + have "sound_ptrs \ bs" + using prems(4) unfolding wf_tree_input_def by blast + hence "predicts (item e)" + using Null prems(6,7) nth_mem entry unfolding sound_ptrs_def sound_null_ptr_def by blast + thus ?thesis + unfolding wf_yield_tree_def predicts_def using simp entry by (auto simp: slice_empty) + next + case Pre + then obtain pre where pre: "pointer e = Pre pre" + by blast + obtain N ts where Nts: "build_tree' bs \ (k-1) pre = Some (Branch N ts)" + using build_tree'_termination entry pre prems(4) wf_tree_input_pre by blast + have simp: "build_tree' bs \ k i = Some (Branch N (ts @ [Leaf (\!(k-1))]))" + using build_tree'_simps(3) entry pre Nts by simp + have "sound_ptrs \ bs" + using prems(4) unfolding wf_tree_input_def by blast + hence bounds: "k > 0" "pre < length (bs!(k-1))" + using entry pre prems(6,7) unfolding sound_ptrs_def sound_pre_ptr_def by (metis nth_mem)+ + moreover have "k-1 < length bs" + by (simp add: prems(6) less_imp_diff_less) + ultimately have IH: "wf_yield_tree \ (item (bs!(k-1)!pre)) (Branch N ts)" + using prems(1) entry pre Nts wf_tree_input_pre prems(4,5,8) by fastforce + have scans: "scans \ k (item (bs!(k-1)!pre)) (item e)" + using entry pre prems(6-7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_pre_ptr_def by simp + have wf: + "item_origin (item (bs!(k-1)!pre)) \ item_end (item (bs!(k-1)!pre))" + "item_end (item (bs!(k-1)!pre)) = k-1" + "item_end (item e) = k" + using entry prems(5,6,7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def wf_item_def + by (auto, meson less_imp_diff_less nth_mem) + have "yield_tree (Branch N (ts @ [Leaf (\!(k-1))])) = concat (map yield_tree (ts @ [Leaf (\!(k-1))]))" + by simp + also have "... = concat (map yield_tree ts) @ [\!(k-1)]" + by simp + also have "... = slice (item_origin (item (bs!(k-1)!pre))) (item_end (item (bs!(k-1)!pre))) \ @ [\!(k-1)]" + using IH by (simp add: wf_yield_tree_def) + also have "... = slice (item_origin (item (bs!(k-1)!pre))) (item_end (item (bs!(k-1)!pre)) + 1) \" + using slice_append_nth wf \k > 0\ prems(8) + by (metis diff_less le_eq_less_or_eq less_imp_diff_less less_numeral_extra(1)) + also have "... = slice (item_origin (item e)) (item_end (item (bs!(k-1)!pre)) + 1) \" + using scans unfolding scans_def inc_item_def by simp + also have "... = slice (item_origin (item e)) k \" + using scans wf unfolding scans_def by (metis Suc_diff_1 Suc_eq_plus1 bounds(1)) + also have "... = slice (item_origin (item e)) (item_end (item e)) \" + using wf by auto + finally show ?thesis + using wf_yield_tree_def entry prems(9) simp by force + next + case PreRed + then obtain k' pre red ps where prered: "pointer e = PreRed (k', pre, red) ps" + by blast + obtain N ts where Nts: "build_tree' bs \ k' pre = Some (Branch N ts)" + using build_tree'_termination entry prems(4) prered wf_tree_input_prered_pre by blast + obtain N_red ts_red where Nts_red: "build_tree' bs \ k red = Some (Branch N_red ts_red)" + using build_tree'_termination entry prems(4) prered wf_tree_input_prered_red by blast + have simp: "build_tree' bs \ k i = Some (Branch N (ts @ [Branch N_red ts_red]))" + using build_tree'_simps(8) entry prered Nts Nts_red by auto + have "sound_ptrs \ bs" + using prems(4) wf_tree_input_def by fastforce + have bounds: "k' < k" "pre < length (bs!k')" "red < length (bs!k)" + using prered entry prems(6,7) \sound_ptrs \ bs\ + unfolding sound_ptrs_def sound_prered_ptr_def by fastforce+ + have completes: "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" + using prered entry prems(6,7) \sound_ptrs \ bs\ + unfolding sound_ptrs_def sound_prered_ptr_def by fastforce + have "(bs, \, k', pre) \ wf_tree_input" + using wf_tree_input_prered_pre[OF prems(4) entry prered] by blast + hence IH_pre: "wf_yield_tree \ (item (bs!k'!pre)) (Branch N ts)" + using prems(2)[OF entry prered _ prems(5)] Nts bounds(1,2) prems(6-8) + by (meson dual_order.strict_trans1 nat_less_le) + have "(bs, \, k, red) \ wf_tree_input" + using wf_tree_input_prered_red[OF prems(4) entry prered] by blast + hence IH_r: "wf_yield_tree \ (item (bs!k!red)) (Branch N_red ts_red)" + using bounds(3) entry prems(3,5,6,8) prered Nts_red by blast + have wf1: + "item_origin (item (bs!k'!pre)) \ item_end (item (bs!k'!pre))" + "item_origin (item (bs!k!red)) \ item_end (item (bs!k!red))" + using prems(5-7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def wf_item_def + by (metis length_map nth_map nth_mem order_less_trans)+ + have wf2: + "item_end (item (bs!k!red)) = k" + "item_end (item (bs!k!i)) = k" + using prems(5-7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def by simp_all + have "yield_tree (Branch N (ts @ [Branch N_red ts_red])) = concat (map yield_tree (ts @ [Branch N_red ts_red]))" + by (simp add: Nts_red) + also have "... = concat (map yield_tree ts) @ yield_tree (Branch N_red ts_red)" + by simp + also have "... = slice (item_origin (item (bs!k'!pre))) (item_end (item (bs!k'!pre))) \ @ + slice (item_origin (item (bs!k!red))) (item_end (item (bs!k!red))) \" + using IH_pre IH_r by (simp add: wf_yield_tree_def) + also have "... = slice (item_origin (item (bs!k'!pre))) (item_end (item (bs!k!red))) \" + using slice_concat wf1 completes_def completes by (metis (no_types, lifting)) + also have "... = slice (item_origin (item e)) (item_end (item (bs!k!red))) \" + using completes unfolding completes_def inc_item_def by simp + also have "... = slice (item_origin (item e)) (item_end (item e)) \" + using wf2 entry by presburger + finally show ?thesis + using wf_yield_tree_def entry prems(9) simp by force + qed + qed + done + thus ?thesis + using assms(2) by blast +qed + +theorem wf_rule_root_yield_tree_build_forest: + assumes "wf_bins \ \ bs" "sound_ptrs \ bs" "mono_red_ptr bs" "length bs = length \ + 1" + assumes "build_tree \ \ bs = Some t" + shows "wf_rule_tree \ t \ root_tree t = \ \ \ yield_tree t = \" +proof - + let ?k = "length bs - 1" + define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items (bs!?k))" + then obtain x i where *: "(x,i) \ set finished" "Some t = build_tree' bs \ ?k i" + using assms(5) unfolding finished_def build_tree_def by (auto simp: Let_def split: list.splits, presburger) + have k: "?k < length bs" "?k \ length \" + using assms(4) by simp_all + have i: "i < length (bs!?k)" + using index_filter_with_index_lt_length * items_def finished_def by (metis length_map) + have x: "x = item (bs!?k!i)" + using * i filter_with_index_nth items_def nth_map finished_def by metis + have finished: "is_finished \ \ x" + using * filter_with_index_P finished_def by metis + have wf_trees_input: "(bs, \, ?k, i) \ wf_tree_input" + unfolding wf_tree_input_def using assms(2,3) i k(1) by blast + hence wf_item_tree: "wf_item_tree \ x t" + using wf_item_tree_build_tree' assms(1,2) i k(1) x *(2) by metis + have wf_item: "wf_item \ \ (item (bs!?k!i))" + using k(1) i assms(1) unfolding wf_bins_def wf_bin_def wf_bin_items_def by (simp add: items_def) + obtain N ts where t: "t = Branch N ts" + by (metis *(2) build_tree'_termination option.inject wf_trees_input) + hence "N = item_rule_head x" + "map root_tree ts = item_rule_body x" + using finished wf_item_tree by (auto simp: is_finished_def is_complete_def) + hence "\r \ set (\ \). N = rule_head r \ map root_tree ts = rule_body r" + using wf_item x unfolding wf_item_def item_rule_body_def item_rule_head_def by blast + hence wf_rule: "wf_rule_tree \ t" + using wf_item_tree t by simp + have root: "root_tree t = \ \" + using finished t \N = item_rule_head x\ by (auto simp: is_finished_def) + have "yield_tree t = slice (item_origin (item (bs!?k!i))) (item_end (item (bs!?k!i))) \" + using k i assms(1) wf_trees_input wf_yield_tree_build_tree' wf_yield_tree_def *(2) by (metis (no_types, lifting)) + hence yield: "yield_tree t = \" + using finished x unfolding is_finished_def by simp + show ?thesis + using * wf_rule root yield assms(4) unfolding build_tree_def by simp +qed + +corollary wf_rule_root_yield_tree_build_tree_Earley\<^sub>L: + assumes "wf_\ \" "nonempty_derives \" + assumes "build_tree \ \ (Earley\<^sub>L \ \) = Some t" + shows "wf_rule_tree \ t \ root_tree t = \ \ \ yield_tree t = \" + using assms wf_rule_root_yield_tree_build_forest wf_bins_Earley\<^sub>L sound_mono_ptrs_Earley\<^sub>L Earley\<^sub>L_def + length_Earley\<^sub>L_bins length_bins_Init\<^sub>L by (metis nle_le) + +theorem correctness_build_tree_Earley\<^sub>L: + assumes "wf_\ \" "is_word \ \" "nonempty_derives \" + shows "(\t. build_tree \ \ (Earley\<^sub>L \ \) = Some t) \ derives \ [\ \] \" (is "?L \ ?R") +proof standard + assume *: ?L + let ?k = "length (Earley\<^sub>L \ \) - 1" + define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items ((Earley\<^sub>L \ \)!?k))" + then obtain t x i where *: "(x,i) \ set finished" "Some t = build_tree' (Earley\<^sub>L \ \) \ ?k i" + using * unfolding finished_def build_tree_def by (auto simp: Let_def split: list.splits, presburger) + have k: "?k < length (Earley\<^sub>L \ \)" "?k \ length \" + by (simp_all add: Earley\<^sub>L_def assms(1)) + have i: "i < length ((Earley\<^sub>L \ \) ! ?k)" + using index_filter_with_index_lt_length * items_def finished_def by (metis length_map) + have x: "x = item ((Earley\<^sub>L \ \)!?k!i)" + using * i filter_with_index_nth items_def nth_map finished_def by metis + have finished: "is_finished \ \ x" + using * filter_with_index_P finished_def by metis + moreover have "x \ set (items ((Earley\<^sub>L \ \) ! ?k))" + using x by (auto simp: items_def; metis One_nat_def i imageI nth_mem) + ultimately have "recognizing (bins (Earley\<^sub>L \ \)) \ \" + by (meson k(1) kth_bin_sub_bins recognizing_def subsetD) + thus ?R + using correctness_Earley\<^sub>L assms by blast +next + assume *: ?R + let ?k = "length (Earley\<^sub>L \ \) - 1" + define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items ((Earley\<^sub>L \ \)!?k))" + have "recognizing (bins (Earley\<^sub>L \ \)) \ \" + using assms * correctness_Earley\<^sub>L by blast + moreover have "?k = length \" + by (simp add: Earley\<^sub>L_def assms(1)) + ultimately have "\x \ set (items ((Earley\<^sub>L \ \)!?k)). is_finished \ \ x" + unfolding recognizing_def using assms(1) is_finished_def wf_bins_Earley\<^sub>L wf_item_in_kth_bin by metis + then obtain x i xs where xis: "finished = (x,i)#xs" + using filter_with_index_Ex_first by (metis finished_def) + hence simp: "build_tree \ \ (Earley\<^sub>L \ \) = build_tree' (Earley\<^sub>L \ \) \ ?k i" + unfolding build_tree_def finished_def by auto + have "(x,i) \ set finished" + using xis by simp + hence "i < length ((Earley\<^sub>L \ \)!?k)" + using index_filter_with_index_lt_length by (metis finished_def items_def length_map) + moreover have "?k < length (Earley\<^sub>L \ \)" + by (simp add: Earley\<^sub>L_def assms(1)) + ultimately have "(Earley\<^sub>L \ \, \, ?k, i) \ wf_tree_input" + unfolding wf_tree_input_def using sound_mono_ptrs_Earley\<^sub>L assms(1,3) by blast + then obtain N ts where "build_tree' (Earley\<^sub>L \ \) \ ?k i = Some (Branch N ts)" + using build_tree'_termination by blast + thus ?L + using simp by simp +qed + + +subsection \those, map, map option lemmas\ + +lemma those_map_exists: + "Some ys = those (map f xs) \ y \ set ys \ \x. x \ set xs \ Some y \ set (map f xs)" +proof (induction xs arbitrary: ys) + case (Cons a xs) + then show ?case + apply (clarsimp split: option.splits) + by (smt (verit, best) map_option_eq_Some set_ConsD) +qed auto + +lemma those_Some: + "(\x \ set xs. \a. x = Some a) \ (\ys. those xs = Some ys)" + by (induction xs) (auto split: option.splits) + +lemma those_Some_P: + assumes "\x \ set xs. \ys. x = Some ys \ (\y \ set ys. P y)" + shows "\yss. those xs = Some yss \ (\ys \ set yss. \y \ set ys. P y)" + using assms by (induction xs) auto + +lemma map_Some_P: + assumes "z \ set (map f xs)" + assumes "\x \ set xs. \ys. f x = Some ys \ (\y \ set ys. P y)" + shows "\ys. z = Some ys \ (\y \ set ys. P y)" + using assms by (induction xs) auto + +lemma those_map_FBranch_only: + assumes "g = (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]])) | _ \ None)" + assumes "Some fs = those (map g pres)" "f \ set fs" + assumes "\f \ set pres. \N fss. f = FBranch N fss" + shows "\f_pre N fss. f = FBranch N (fss @ [[FLeaf (\!(k-1))]]) \ f_pre = FBranch N fss \ f_pre \ set pres" + using assms + apply (induction pres arbitrary: fs f) + apply (auto) + by (smt (verit, best) list.inject list.set_cases map_option_eq_Some) + +lemma those_map_Some_concat_exists: + assumes "y \ set (concat ys)" + assumes "Some ys = those (map f xs)" + shows "\ys x. Some ys = f x \ y \ set ys \ x \ set xs" + using assms + apply (induction xs arbitrary: ys y) + apply (auto split: option.splits) + by (smt (verit, ccfv_threshold) list.inject list.set_cases map_option_eq_Some) + +lemma map_option_concat_those_map_exists: + assumes "Some fs = map_option concat (those (map F xs))" + assumes "f \ set fs" + shows "\fss fs'. Some fss = those (map F xs) \ fs' \ set fss \ f \ set fs'" + using assms + apply (induction xs arbitrary: fs f) + apply (auto split: option.splits) + by (smt (verit, best) UN_E map_option_eq_Some set_concat) + +lemma [partial_function_mono]: + "monotone option.le_fun option_ord + (\f. map_option concat (those (map (\((k', pre), reds). + f ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. f ((((r, s), t), red), b \ {red})) reds) \ + (\rss. those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None) pres)))) + xs)))" +proof - + let ?f = " + (\f. map_option concat (those (map (\((k', pre), reds). + f ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. f ((((r, s), t), red), b \ {red})) reds) \ + (\rss. those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None) pres)))) + xs)))" + have 0: "\x y. option.le_fun x y \ option_ord (?f x) (?f y)" + apply (auto simp: flat_ord_def fun_ord_def option.leq_refl split: option.splits forest.splits) + subgoal premises prems for x y + proof - + let ?t = "those (map (\((k', pre), reds). + x ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))) + xs) = None" + show ?t + proof (rule ccontr) + assume a: "\?t" + obtain fss where fss: "those (map (\((k', pre), reds). + x ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))) + xs) = Some fss" + using a by blast + { + fix k' pre reds + assume *: "((k', pre), reds) \ set xs" + obtain pres where pres: "x ((((r, s), k'), pre), {pre}) = Some pres" + using fss * those_Some by force + have "\fs. Some fs = those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))" + proof (rule ccontr) + assume "\fs. Some fs = + those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))" + hence "None = + those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))" + by (smt (verit) not_None_eq) + hence "None = x ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))" + by (simp add: pres) + hence "\((k', pre), reds) \ set xs. None = x ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))" + using * by blast + thus False + using fss those_Some by force + qed + then obtain fs where fs: "Some fs = those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))" + by blast + obtain rss where rss: "those (map (\red. x ((((r, s), t), red), insert red b)) reds) = Some rss" + using fs by force + have "x ((((r, s), k'), pre), {pre}) = y ((((r, s), k'), pre), {pre})" + using pres prems(1) by (metis option.distinct(1)) + moreover have "those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)) + = those (map (\red. y ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))" + proof - + have "\red \ set reds. x ((((r, s), t), red), insert red b) = y ((((r, s), t), red), insert red b)" + proof standard + fix red + assume "red \ set reds" + have "\x\set (map (\red. x ((((r, s), t), red), insert red b)) reds) . \a. x = Some a" + using rss those_Some by blast + then obtain f where "x ((((r, s), t), red), insert red b) = Some f" + using \red \ set reds\ by auto + thus "x ((((r, s), t), red), insert red b) = y ((((r, s), t), red), insert red b)" + using prems(1) by (metis option.distinct(1)) + qed + thus ?thesis + by (smt (verit, best) map_eq_conv) + qed + ultimately have " x ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))) + = y ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. y ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))" + by (metis bind.bind_lunit pres) + } + hence "\((k', pre), reds) \ set xs. x ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres))) + = y ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. y ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))" + by blast + hence "those (map (\((k', pre), reds). + x ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. x ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))) + xs) = those (map (\((k', pre), reds). + y ((((r, s), k'), pre), {pre}) \ + (\pres. those (map (\red. y ((((r, s), t), red), insert red b)) reds) \ + (\rss. those (map (case_forest Map.empty (\N fss. Some (FBranch N (fss @ [concat rss])))) pres)))) + xs)" + using prems(1) by (smt (verit, best) case_prod_conv map_eq_conv split_cong) + thus False + using prems(2) by simp + qed + qed + done + show ?thesis + using monotoneI[of option.le_fun option_ord ?f, OF 0] by blast +qed + +subsection \Parse trees\ + +fun insert_group :: "('a \ 'k) \ ('a \ 'v) \ 'a \ ('k \ 'v list) list \ ('k \ 'v list) list" where + "insert_group K V a [] = [(K a, [V a])]" +| "insert_group K V a ((k, vs)#xs) = ( + if K a = k then (k, V a # vs) # xs + else (k, vs) # insert_group K V a xs + )" + +fun group_by :: "('a \ 'k) \ ('a \ 'v) \ 'a list \ ('k \ 'v list) list" where + "group_by K V [] = []" +| "group_by K V (x#xs) = insert_group K V x (group_by K V xs)" + +lemma insert_group_cases: + assumes "(k, vs) \ set (insert_group K V a xs)" + shows "(k = K a \ vs = [V a]) \ (k, vs) \ set xs \ (\(k', vs') \ set xs. k' = k \ k = K a \ vs = V a # vs')" + using assms by (induction xs) (auto split: if_splits) + +lemma group_by_exists_kv: + "(k, vs) \ set (group_by K V xs) \ \x \ set xs. k = K x \ (\v \ set vs. v = V x)" + using insert_group_cases by (induction xs) (simp, force) + +lemma group_by_forall_v_exists_k: + "(k, vs) \ set (group_by K V xs) \ v \ set vs \ \x \ set xs. k = K x \ v = V x" +proof (induction xs arbitrary: vs) + case (Cons x xs) + show ?case + proof (cases "(k, vs) \ set (group_by K V xs)") + case True + thus ?thesis + using Cons by simp + next + case False + hence "(k, vs) \ set (insert_group K V x (group_by K V xs))" + using Cons.prems(1) by force + then consider (A) "(k = K x \ vs = [V x])" + | (B) "(k, vs) \ set (group_by K V xs)" + | (C) "(\(k', vs') \ set (group_by K V xs). k' = k \ k = K x \ vs = V x # vs')" + using insert_group_cases by fastforce + thus ?thesis + proof cases + case A + thus ?thesis + using Cons.prems(2) by auto + next + case B + then show ?thesis + using False by linarith + next + case C + then show ?thesis + using Cons.IH Cons.prems(2) by fastforce + qed + qed +qed simp + +partial_function (option) build_trees' :: "'a bins \ 'a sentence \ nat \ nat \ nat set \ 'a forest list option" where + "build_trees' bs \ k i I = ( + let e = bs!k!i in ( + case pointer e of + Null \ Some ([FBranch (item_rule_head (item e)) []]) \\start building sub-trees\ + | Pre pre \ ( \\add sub-trees starting from terminal\ + do { + pres \ build_trees' bs \ (k-1) pre {pre}; + those (map (\f. + case f of + FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]])) + | _ \ None \\impossible case\ + ) pres) + }) + | PreRed p ps \ ( \\add sub-trees starting from non-terminal\ + let ps' = filter (\(k', pre, red). red \ I) (p#ps) in + let gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps' in + map_option concat (those (map (\((k', pre), reds). + do { + pres \ build_trees' bs \ k' pre {pre}; + rss \ those (map (\red. build_trees' bs \ k red (I \ {red})) reds); + those (map (\f. + case f of + FBranch N fss \ Some (FBranch N (fss @ [concat rss])) + | _ \ None \\impossible case\ + ) pres) + } + ) gs)) + ) + ))" + +declare build_trees'.simps [code] + +definition build_trees :: "'a cfg \ 'a sentence \ 'a bins \ 'a forest list option" where + "build_trees \ \ bs = ( + let k = length bs - 1 in + let finished = filter_with_index (\x. is_finished \ \ x) (items (bs!k)) in + map_option concat (those (map (\(_, i). build_trees' bs \ k i {i}) finished)) + )" + +lemma build_forest'_simps[simp]: + "e = bs!k!i \ pointer e = Null \ build_trees' bs \ k i I = Some ([FBranch (item_rule_head (item e)) []])" + "e = bs!k!i \ pointer e = Pre pre \ build_trees' bs \ (k-1) pre {pre} = None \ build_trees' bs \ k i I = None" + "e = bs!k!i \ pointer e = Pre pre \ build_trees' bs \ (k-1) pre {pre} = Some pres \ + build_trees' bs \ k i I = those (map (\f. case f of FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]])) | _ \ None) pres)" + by (subst build_trees'.simps, simp)+ + +definition wf_trees_input :: "('a bins \ 'a sentence \ nat \ nat \ nat set) set" where + "wf_trees_input = { + (bs, \, k, i, I) | bs \ k i I. + sound_ptrs \ bs \ + k < length bs \ + i < length (bs!k) \ + I \ {0.. + i \ I + }" + +fun build_forest'_measure :: "('a bins \ 'a sentence \ nat \ nat \ nat set) \ nat" where + "build_forest'_measure (bs, \, k, i, I) = foldl (+) 0 (map length (take (k+1) bs)) - card I" + +lemma wf_trees_input_pre: + assumes "(bs, \, k, i, I) \ wf_trees_input" + assumes "e = bs!k!i" "pointer e = Pre pre" + shows "(bs, \, (k-1), pre, {pre}) \ wf_trees_input" + using assms unfolding wf_trees_input_def + apply (auto simp: sound_ptrs_def sound_pre_ptr_def) + apply (metis nth_mem) + done + +lemma wf_trees_input_prered_pre: + assumes "(bs, \, k, i, I) \ wf_trees_input" + assumes "e = bs!k!i" "pointer e = PreRed p ps" + assumes "ps' = filter (\(k', pre, red). red \ I) (p#ps)" + assumes "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" + assumes "((k', pre), reds) \ set gs" + shows "(bs, \, k', pre, {pre}) \ wf_trees_input" +proof - + obtain red where "(k', pre, red) \ set ps'" + using assms(5,6) group_by_exists_kv by fast + hence *: "(k', pre, red) \ set (p#ps)" + using assms(4) by (meson filter_is_subset in_mono) + have "k < length bs" "e \ set (bs!k)" + using assms(1,2) unfolding wf_trees_input_def by auto + hence "k' < k" "pre < length (bs!k')" + using assms(1,3) * unfolding wf_trees_input_def sound_ptrs_def sound_prered_ptr_def by blast+ + thus ?thesis + using assms by (simp add: wf_trees_input_def) +qed + +lemma wf_trees_input_prered_red: + assumes "(bs, \, k, i, I) \ wf_trees_input" + assumes "e = bs!k!i" "pointer e = PreRed p ps" + assumes "ps' = filter (\(k', pre, red). red \ I) (p#ps)" + assumes "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" + assumes "((k', pre), reds) \ set gs" "red \ set reds" + shows "(bs, \, k, red, I \ {red}) \ wf_trees_input" +proof - + have "(k', pre, red) \ set ps'" + using assms(5,6,7) group_by_forall_v_exists_k by fastforce + hence *: "(k', pre, red) \ set (p#ps)" + using assms(4) by (meson filter_is_subset in_mono) + have "k < length bs" "e \ set (bs!k)" + using assms(1,2) unfolding wf_trees_input_def by auto + hence 0: "red < length (bs!k)" + using assms(1,3) * unfolding wf_trees_input_def sound_ptrs_def sound_prered_ptr_def by blast + moreover have "I \ {0.. {red} \ {0.., k, i, I) \ wf_trees_input" + assumes "\bs \ k i I. + (\e pre. e = bs!k!i \ pointer e = Pre pre \ P bs \ (k-1) pre {pre}) \ + (\e p ps ps' gs k' pre reds. e = bs!k!i \ pointer e = PreRed p ps \ + ps' = filter (\(k', pre, red). red \ I) (p#ps) \ + gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps' \ + ((k', pre), reds) \ set gs \ P bs \ k' pre {pre}) \ + (\e p ps ps' gs k' pre red reds reds'. e = bs!k!i \ pointer e = PreRed p ps \ + ps' = filter (\(k', pre, red). red \ I) (p#ps) \ + gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps' \ + ((k', pre), reds) \ set gs \ red \ set reds \ P bs \ k red (I \ {red})) \ + P bs \ k i I" + shows "P bs \ k i I" + using assms(1) +proof (induction n\"build_forest'_measure (bs, \, k, i, I)" arbitrary: k i I rule: nat_less_induct) + case 1 + obtain e where entry: "e = bs!k!i" + by simp + consider (Null) "pointer e = Null" + | (Pre) "\pre. pointer e = Pre pre" + | (PreRed) "\p ps. pointer e = PreRed p ps" + by (metis pointer.exhaust) + thus ?case + proof cases + case Null + thus ?thesis + using assms(2) entry by fastforce + next + case Pre + then obtain pre where pre: "pointer e = Pre pre" + by blast + define n where n: "n = build_forest'_measure (bs, \, (k-1), pre, {pre})" + have "0 < k" "pre < length (bs!(k-1))" + using 1(2) entry pre unfolding wf_trees_input_def sound_ptrs_def sound_pre_ptr_def + by (smt (verit) mem_Collect_eq nth_mem prod.inject)+ + have "k < length bs" "i < length (bs!k)" "I \ {0.. I" + using 1(2) unfolding wf_trees_input_def by blast+ + have "length (bs!(k-1)) > 0" + using \pre < length (bs!(k-1))\ by force + hence "foldl (+) 0 (map length (take k bs)) > 0" + by (smt (verit, del_insts) foldl_add_nth \0 < k\ \k < length bs\ + add.commute add_diff_inverse_nat less_imp_diff_less less_one zero_eq_add_iff_both_eq_0) + have "card I \ length (bs!k)" + by (simp add: \I \ {0.. subset_eq_atLeast0_lessThan_card) + have "card I + (foldl (+) 0 (map length (take (Suc (k - Suc 0)) bs)) - Suc 0) = + card I + (foldl (+) 0 (map length (take k bs)) - 1)" + using \0 < k\ by simp + also have "... = card I + foldl (+) 0 (map length (take k bs)) - 1" + using \0 < foldl (+) 0 (map length (take k bs))\ by auto + also have "... < card I + foldl (+) 0 (map length (take k bs))" + by (simp add: \0 < foldl (+) 0 (map length (take k bs))\) + also have "... \ foldl (+) 0 (map length (take k bs)) + length (bs!k)" + by (simp add: \card I \ length (bs ! k)\) + also have "... = foldl (+) 0 (map length (take (k+1) bs))" + using foldl_add_nth \k < length bs\ by blast + finally have "build_forest'_measure (bs, \, k, i, I) - build_forest'_measure (bs, \, (k-1), pre, {pre}) > 0" + by simp + hence "P bs \ (k-1) pre {pre}" + using 1 n wf_trees_input_pre entry pre zero_less_diff by blast + thus ?thesis + using assms(2) entry pre pointer.distinct(5) pointer.inject(1) by presburger + next + case PreRed + then obtain p ps where pps: "pointer e = PreRed p ps" + by blast + define ps' where ps': "ps' = filter (\(k', pre, red). red \ I) (p#ps)" + define gs where gs: "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" + have 0: "\(k', pre, red) \ set ps'. k' < k \ pre < length (bs!k') \ red < length (bs!k)" + using entry pps ps' 1(2) unfolding wf_trees_input_def sound_ptrs_def sound_prered_ptr_def + apply (auto simp del: filter.simps) + apply (metis nth_mem prod_cases3)+ + done + hence sound_gs: "\((k', pre), reds) \ set gs. k' < k \ pre < length (bs!k')" + using gs group_by_exists_kv by fast + have sound_gs2: "\((k', pre), reds) \ set gs. \red \ set reds. red < length (bs!k)" + proof (standard, standard, standard, standard) + fix x a b k' pre red + assume "x \ set gs" "x = (a, b)" "(k', pre) = a" "red \ set b" + hence "\x \ set ps'. red = (\(k', pre, red). red) x" + using group_by_forall_v_exists_k gs ps' by meson + thus "red < length (bs!k)" + using 0 by fast + qed + { + fix k' pre reds red + assume a0: "((k', pre), reds) \ set gs" + define n_pre where n_pre: "n_pre = build_forest'_measure (bs, \, k', pre, {pre})" + have "k < length bs" "i < length (bs!k)" "I \ {0.. I" + using 1(2) unfolding wf_trees_input_def by blast+ + have "k' < k" "pre < length (bs!k')" + using sound_gs a0 by fastforce+ + have "length (bs!k') > 0" + using \pre < length (bs!k')\ by force + hence gt0: "foldl (+) 0 (map length (take (k'+1) bs)) > 0" + by (smt (verit, del_insts) foldl_add_nth \k < length bs\ \k' < k\ add_gr_0 order.strict_trans) + have card_bound: "card I \ length (bs!k)" + by (simp add: \I \ {0.. subset_eq_atLeast0_lessThan_card) + have "card I + (foldl (+) 0 (map length (take (Suc k') bs)) - Suc 0) = + card I + foldl (+) 0 (map length (take (Suc k') bs)) - 1" + by (metis Nat.add_diff_assoc One_nat_def Suc_eq_plus1 Suc_leI \0 < foldl (+) 0 (map length (take (k' + 1) bs))\) + also have "... < card I + foldl (+) 0 (map length (take (Suc k') bs))" + using gt0 by auto + also have "... \ foldl (+) 0 (map length (take (Suc k') bs)) + length (bs!k)" + using card_bound by simp + also have "... \ foldl (+) 0 (map length (take (k+1) bs))" + using foldl_add_nth_ge Suc_leI \k < length bs\ \k' < k\ by blast + finally have "build_forest'_measure (bs, \, k, i, I) - build_forest'_measure (bs, \, k', pre, {pre}) > 0" + by simp + hence "P bs \ k' pre {pre}" + using 1(1) wf_trees_input_prered_pre[OF "1.prems"(1) entry pps ps' gs a0] zero_less_diff by blast + } + moreover { + fix k' pre reds red + assume a0: "((k', pre), reds) \ set gs" + assume a1: "red \ set reds" + define n_red where n_red: "n_red = build_forest'_measure (bs, \, k, red, I \ {red})" + have "k < length bs" "i < length (bs!k)" "I \ {0.. I" + using 1(2) unfolding wf_trees_input_def by blast+ + have "k' < k" "pre < length (bs!k')" "red < length (bs!k)" + using a0 a1 sound_gs sound_gs2 by fastforce+ + have "red \ I" + using a0 a1 unfolding gs ps' + by (smt (verit, best) group_by_forall_v_exists_k case_prodE case_prod_conv mem_Collect_eq set_filter) + have card_bound: "card I \ length (bs!k)" + by (simp add: \I \ {0.. subset_eq_atLeast0_lessThan_card) + have "length (bs!k') > 0" + using \pre < length (bs!k')\ by force + hence gt0: "foldl (+) 0 (map length (take (k'+1) bs)) > 0" + by (smt (verit, del_insts) foldl_add_nth \k < length bs\ \k' < k\ add_gr_0 order.strict_trans) + have lt: "foldl (+) 0 (map length (take (Suc k') bs)) + length (bs!k) \ foldl (+) 0 (map length (take (k+1) bs))" + using foldl_add_nth_ge Suc_leI \k < length bs\ \k' < k\ by blast + have "card I + (foldl (+) 0 (map length (take (Suc k) bs)) - card (insert red I)) = + card I + (foldl (+) 0 (map length (take (Suc k) bs)) - card I - 1)" + using \I \ {0.. \red \ I\ finite_subset by fastforce + also have "... = foldl (+) 0 (map length (take (Suc k) bs)) - 1" + using gt0 card_bound lt by force + also have "... < foldl (+) 0 (map length (take (Suc k) bs))" + using gt0 lt by auto + finally have "build_forest'_measure (bs, \, k, i, I) - build_forest'_measure (bs, \, k, red, I \ {red}) > 0" + by simp + moreover have "(bs, \, k, red, I \ {red}) \ wf_trees_input" + using wf_trees_input_prered_red[OF "1"(2) entry pps ps' gs] a0 a1 by blast + ultimately have "P bs \ k red (I \ {red})" + using 1(1) zero_less_diff by blast + } + moreover have "(\e pre. e = bs!k!i \ pointer e = Pre pre \ P bs \ (k-1) pre {pre})" + using entry pps by fastforce + ultimately show ?thesis + using assms(2) entry gs pointer.inject(2) pps ps' by presburger + qed +qed + +lemma build_trees'_termination: + assumes "(bs, \, k, i, I) \ wf_trees_input" + shows "\fs. build_trees' bs \ k i I = Some fs \ (\f \ set fs. \N fss. f = FBranch N fss)" +proof - + have "\fs. build_trees' bs \ k i I = Some fs \ (\f \ set fs. \N fss. f = FBranch N fss)" + apply (induction rule: build_trees'_induct[OF assms(1)]) + subgoal premises IH for bs \ k i I + proof - + define e where entry: "e = bs!k!i" + consider (Null) "pointer e = Null" + | (Pre) "\pre. pointer e = Pre pre" + | (PreRed) "\k' pre red reds. pointer e = PreRed (k', pre, red) reds" + by (metis pointer.exhaust surj_pair) + thus ?thesis + proof cases + case Null + have "build_trees' bs \ k i I = Some ([FBranch (item_rule_head (item e)) []])" + using build_forest'_simps(1) Null entry by simp + thus ?thesis + by simp + next + case Pre + then obtain pre where pre: "pointer e = Pre pre" + by blast + obtain fs where fs: "build_trees' bs \ (k-1) pre {pre} = Some fs" + "\f \ set fs. \N fss. f = FBranch N fss" + using IH(1) entry pre by blast + let ?g = "\f. case f of FLeaf a \ None + | FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]]))" + have simp: "build_trees' bs \ k i I = those (map ?g fs)" + using build_forest'_simps(3) entry pre fs by blast + moreover have "\f \ set (map ?g fs). \a. f = Some a" + using fs(2) by auto + ultimately obtain fs' where fs': "build_trees' bs \ k i I = Some fs'" + using those_Some by (smt (verit, best)) + moreover have "\f \ set fs'. \N fss. f = FBranch N fss" + proof standard + fix f + assume "f \ set fs'" + then obtain x where "x \ set fs" "Some f \ set (map ?g fs)" + using those_map_exists by (metis (no_types, lifting) fs' simp) + thus "\N fss. f = FBranch N fss" + using fs(2) by auto + qed + ultimately show ?thesis + by blast + next + case PreRed + then obtain p ps where pps: "pointer e = PreRed p ps" + by blast + define ps' where ps': "ps' = filter (\(k', pre, red). red \ I) (p#ps)" + define gs where gs: "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" + let ?g = "\((k', pre), reds). + do { + pres \ build_trees' bs \ k' pre {pre}; + rss \ those (map (\red. build_trees' bs \ k red (I \ {red})) reds); + those (map (\f. + case f of + FBranch N fss \ Some (FBranch N (fss @ [concat rss])) + | _ \ None \\impossible case\ + ) pres) + }" + have simp: "build_trees' bs \ k i I = map_option concat (those (map ?g gs))" + using entry pps ps' gs by (subst build_trees'.simps) (auto simp del: filter.simps) + have "\fso \ set (map ?g gs). \fs. fso = Some fs \ (\f \ set fs. \N fss. f = FBranch N fss)" + proof standard + fix fso + assume "fso \ set (map ?g gs)" + moreover have "\ps \ set gs. \fs. ?g ps = Some fs \ (\f \ set fs. \N fss. f = FBranch N fss)" + proof standard + fix ps + assume "ps \ set gs" + then obtain k' pre reds where *: "((k', pre), reds) \ set gs" "((k', pre), reds) = ps" + by (metis surj_pair) + then obtain pres where pres: "build_trees' bs \ k' pre {pre} = Some pres" + "\f \ set pres. \N fss. f = FBranch N fss" + using IH(2) entry pps ps' gs by blast + have "\f \ set (map (\red. build_trees' bs \ k red (I \ {red})) reds). \a. f = Some a" + using IH(3)[OF entry pps ps' gs *(1)] by auto + then obtain rss where rss: "Some rss = those (map (\red. build_trees' bs \ k red (I \ {red})) reds)" + using those_Some by (metis (full_types)) + let ?h = "\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None" + have "\x \ set (map ?h pres). \a. x = Some a" + using pres(2) by auto + then obtain fs where fs: "Some fs = those (map ?h pres)" + using those_Some by (smt (verit, best)) + have "\f \ set fs. \N fss. f = FBranch N fss" + proof standard + fix f + assume *: "f \ set fs" + hence "\x. x \ set pres \ Some f \ set (map ?h pres)" + using those_map_exists[OF fs *] by blast + then obtain x where x: "x \ set pres \ Some f \ set (map ?h pres)" + by blast + thus "\N fss. f = FBranch N fss" + using pres(2) by auto + qed + moreover have "?g ps = Some fs" + using fs pres rss * by (auto, metis bind.bind_lunit) + ultimately show "\fs. ?g ps = Some fs \ (\f\set fs. \N fss. f = FBranch N fss)" + by blast + qed + ultimately show "\fs. fso = Some fs \ (\f \ set fs. \N fss. f = FBranch N fss)" + using map_Some_P by auto + qed + then obtain fss where "those (map ?g gs) = Some fss" "\fs \ set fss. \f \ set fs. \N fss. f = FBranch N fss" + using those_Some_P by blast + hence "build_trees' bs \ k i I = Some (concat fss)" "\f \ set (concat fss). \N fss. f = FBranch N fss" + using simp by auto + thus ?thesis + by blast + qed + qed + done + thus ?thesis + by blast +qed + +lemma wf_item_tree_build_trees': + assumes "(bs, \, k, i, I) \ wf_trees_input" + assumes "wf_bins \ \ bs" + assumes "k < length bs" "i < length (bs!k)" + assumes "build_trees' bs \ k i I = Some fs" + assumes "f \ set fs" + assumes "t \ set (trees f)" + shows "wf_item_tree \ (item (bs!k!i)) t" +proof - + have "wf_item_tree \ (item (bs!k!i)) t" + using assms + apply (induction arbitrary: fs f t rule: build_trees'_induct[OF assms(1)]) + subgoal premises prems for bs \ k i I fs f t + proof - + define e where entry: "e = bs!k!i" + consider (Null) "pointer e = Null" + | (Pre) "\pre. pointer e = Pre pre" + | (PreRed) "\p ps. pointer e = PreRed p ps" + by (metis pointer.exhaust) + thus ?thesis + proof cases + case Null + hence simp: "build_trees' bs \ k i I = Some ([FBranch (item_rule_head (item e)) []])" + using entry by simp + moreover have "f = FBranch (item_rule_head (item e)) []" + using build_forest'_simps(1) Null prems(8,9) entry by auto + ultimately have simp: "t = Branch (item_rule_head (item e)) []" + using prems(10) by simp + have "sound_ptrs \ bs" + using prems(4) unfolding wf_trees_input_def by blast + hence "predicts (item e)" + using Null prems(6,7) nth_mem entry unfolding sound_ptrs_def sound_null_ptr_def by blast + hence "item_dot (item e) = 0" + unfolding predicts_def by blast + thus ?thesis + using simp entry by simp + next + case Pre + then obtain pre where pre: "pointer e = Pre pre" + by blast + have sound: "sound_ptrs \ bs" + using prems(4) unfolding wf_trees_input_def by blast + have scans: "scans \ k (item (bs!(k-1)!pre)) (item e)" + using entry pre prems(6-7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_pre_ptr_def by simp + hence *: + "item_rule_head (item (bs!(k-1)!pre)) = item_rule_head (item e)" + "item_rule_body (item (bs!(k-1)!pre)) = item_rule_body (item e)" + "item_dot (item (bs!(k-1)!pre)) + 1 = item_dot (item e)" + "next_symbol (item (bs!(k-1)!pre)) = Some (\!(k-1))" + unfolding scans_def inc_item_def by (simp_all add: item_rule_head_def item_rule_body_def) + have wf: "(bs, \, k-1, pre, {pre}) \ wf_trees_input" + using entry pre prems(4) wf_trees_input_pre by blast + then obtain pres where pres: "build_trees' bs \ (k-1) pre {pre} = Some pres" + "\f \ set pres. \N fss. f = FBranch N fss" + using build_trees'_termination wf by blast + let ?g = "\f. case f of FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]])) | _ \ None" + have "build_trees' bs \ k i I = those (map ?g pres)" + using entry pre pres by simp + hence fs: "Some fs = those (map ?g pres)" + using prems(8) by simp + then obtain f_pre N fss where Nfss: "f = FBranch N (fss @ [[FLeaf (\!(k-1))]])" + "f_pre = FBranch N fss" "f_pre \ set pres" + using those_map_FBranch_only fs pres(2) prems(9) by blast + define tss where tss: "tss = map (\fs. concat (map (\f. trees f) fs)) fss" + have "trees (FBranch N (fss @ [[FLeaf (\!(k-1))]])) = + map (\ts. Branch N ts) [ ts @ [Leaf (\!(k-1))] . ts <- combinations tss ]" + by (subst tss, subst trees_append_single_singleton, simp) + moreover have "t \ set (trees (FBranch N (fss @ [[FLeaf (\!(k-1))]])))" + using Nfss(1) prems(10) by blast + ultimately obtain ts where ts: "t = Branch N (ts @ [Leaf (\!(k-1))]) \ ts \ set (combinations tss)" + by auto + have "sound_ptrs \ bs" + using prems(4) unfolding wf_trees_input_def by blast + hence "pre < length (bs!(k-1))" + using entry pre prems(6,7) unfolding sound_ptrs_def sound_pre_ptr_def by (metis nth_mem) + moreover have "k - 1 < length bs" + by (simp add: prems(6) less_imp_diff_less) + moreover have "Branch N ts \ set (trees (FBranch N fss))" + using ts tss by simp + ultimately have IH: "wf_item_tree \ (item (bs!(k-1)!pre)) (Branch N ts)" + using prems(1,2,4,5) entry pre Nfss(2,3) wf pres(1) by blast + have "map root_tree (ts @ [Leaf (\!(k-1))]) = map root_tree ts @ [\!(k-1)]" + by simp + also have "... = take (item_dot (item (bs!(k-1)!pre))) (item_rule_body (item (bs!(k-1)!pre))) @ [\!(k-1)]" + using IH by simp + also have "... = take (item_dot (item (bs!(k-1)!pre))) (item_rule_body (item e)) @ [\!(k-1)]" + using *(2) by simp + also have "... = take (item_dot (item e)) (item_rule_body (item e))" + using *(2-4) by (auto simp: next_symbol_def is_complete_def split: if_splits; metis leI take_Suc_conv_app_nth) + finally have "map root_tree (ts @ [Leaf (\!(k-1))]) = take (item_dot (item e)) (item_rule_body (item e))" . + hence "wf_item_tree \ (item e) (Branch N (ts @ [Leaf (\!(k-1))]))" + using IH *(1) by simp + thus ?thesis + using ts entry by fastforce + next + case PreRed + then obtain p ps where prered: "pointer e = PreRed p ps" + by blast + define ps' where ps': "ps' = filter (\(k', pre, red). red \ I) (p#ps)" + define gs where gs: "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" + let ?g = "\((k', pre), reds). + do { + pres \ build_trees' bs \ k' pre {pre}; + rss \ those (map (\red. build_trees' bs \ k red (I \ {red})) reds); + those (map (\f. + case f of + FBranch N fss \ Some (FBranch N (fss @ [concat rss])) + | _ \ None \\impossible case\ + ) pres) + }" + have simp: "build_trees' bs \ k i I = map_option concat (those (map ?g gs))" + using entry prered ps' gs by (subst build_trees'.simps) (auto simp del: filter.simps) + have "\fso \ set (map ?g gs). \fs. fso = Some fs \ (\f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t)" + proof standard + fix fso + assume "fso \ set (map ?g gs)" + moreover have "\ps \ set gs. \fs. ?g ps = Some fs \ (\f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t)" + proof standard + fix g + assume "g \ set gs" + then obtain k' pre reds where g: "((k', pre), reds) \ set gs" "((k', pre), reds) = g" + by (metis surj_pair) + moreover have wf_pre: "(bs, \, k', pre, {pre}) \ wf_trees_input" + using wf_trees_input_prered_pre[OF prems(4) entry prered ps' gs g(1)] by blast + ultimately obtain pres where pres: "build_trees' bs \ k' pre {pre} = Some pres" + "\f_pre \ set pres. \N fss. f_pre = FBranch N fss" + using build_trees'_termination by blast + have wf_reds: "\red \ set reds. (bs, \, k, red, I \ {red}) \ wf_trees_input" + using wf_trees_input_prered_red[OF prems(4) entry prered ps' gs g(1)] by blast + hence "\f \ set (map (\red. build_trees' bs \ k red (I \ {red})) reds). \a. f = Some a" + using build_trees'_termination by fastforce + then obtain rss where rss: "Some rss = those (map (\red. build_trees' bs \ k red (I \ {red})) reds)" + using those_Some by (metis (full_types)) + let ?h = "\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None" + have "\x \ set (map ?h pres). \a. x = Some a" + using pres(2) by auto + then obtain fs where fs: "Some fs = those (map ?h pres)" + using those_Some by (smt (verit, best)) + have "\f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t" + proof (standard, standard) + fix f t + assume ft: "f \ set fs" "t \ set (trees f)" + hence "\x. x \ set pres \ Some f \ set (map ?h pres)" + using those_map_exists[OF fs ft(1)] by blast + then obtain f_pre N fss where f_pre: "f_pre \ set pres" "f_pre = FBranch N fss" + "f = FBranch N (fss @ [concat rss])" + using pres(2) by force + define tss where tss: "tss = map (\fs. concat (map (\f. trees f) fs)) fss" + have "trees (FBranch N (fss @ [concat rss])) = + map (\ts. Branch N ts) [ ts0 @ ts1 . ts0 <- combinations tss, + ts1 <- combinations [concat (map (\f. trees f) (concat rss)) ] ]" + by (subst tss, subst trees_append_singleton, simp) + moreover have "t \ set (trees (FBranch N (fss @ [concat rss])))" + using ft(2) f_pre(3) by blast + ultimately obtain ts0 ts1 where tsx: "t = Branch N (ts0 @ [ts1])" "ts0 \ set (combinations tss)" + "ts1 \ set (concat (map (\f. trees f) (concat rss)))" + by fastforce + then obtain f_red where f_red: "f_red \ set (concat rss)" "ts1 \ set (trees f_red)" + by auto + obtain fs_red red where red: "Some fs_red = build_trees' bs \ k red (I \ {red})" + "f_red \ set fs_red" "red \ set reds" + using f_red(1) rss those_map_Some_concat_exists by fast + then obtain N_red fss_red where "f_red = FBranch N_red fss_red" + using build_trees'_termination wf_reds by (metis option.inject) + then obtain ts where ts: "Branch N_red ts = ts1" + using tsx(3) f_red by auto + have "(k', pre, red) \ set ps'" + using group_by_forall_v_exists_k \((k', pre), reds) \ set gs\ gs \red \ set reds\ by fast + hence mem: "(k', pre, red) \ set (p#ps)" + using ps' by (metis filter_set member_filter) + have "sound_ptrs \ bs" + using prems(4) wf_trees_input_def by fastforce + have bounds: "k' < k" "pre < length (bs!k')" "red < length (bs!k)" + using prered entry prems(6,7) \sound_ptrs \ bs\ + unfolding sound_ptrs_def sound_prered_ptr_def by (meson mem nth_mem)+ + have completes: "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" + using prered entry prems(6,7) \sound_ptrs \ bs\ + unfolding sound_ptrs_def sound_prered_ptr_def by (metis mem nth_mem) + have transform: + "item_rule_head (item (bs!k'!pre)) = item_rule_head (item e)" + "item_rule_body (item (bs!k'!pre)) = item_rule_body (item e)" + "item_dot (item (bs!k'!pre)) + 1 = item_dot (item e)" + "next_symbol (item (bs!k'!pre)) = Some (item_rule_head (item (bs!k!red)))" + "is_complete (item (bs!k!red))" + using completes unfolding completes_def inc_item_def + by (auto simp: item_rule_head_def item_rule_body_def is_complete_def) + have "Branch N ts0 \ set (trees (FBranch N fss))" + using tss tsx(2) by simp + hence IH_pre: "wf_item_tree \ (item (bs!k'!pre)) (Branch N ts0)" + using prems(2)[OF entry prered ps' gs \((k', pre), reds) \ set gs\ wf_pre prems(5)] + pres(1) f_pre f_pre(3) bounds(1,2) prems(6) by fastforce + have IH_r: "wf_item_tree \ (item (bs!k!red)) (Branch N_red ts)" + using prems(3)[OF entry prered ps' gs \((k', pre), reds) \ set gs\ \red \ set reds\ _ prems(5)] + bounds(3) f_red(2) red ts wf_reds prems(6) by metis + have "map root_tree (ts0 @ [Branch N_red ts]) = map root_tree ts0 @ [root_tree (Branch N_red ts)]" + by simp + also have "... = take (item_dot (item (bs!k'!pre))) (item_rule_body (item (bs!k'!pre))) @ [root_tree (Branch N_red ts)]" + using IH_pre by simp + also have "... = take (item_dot (item (bs!k'!pre))) (item_rule_body (item (bs!k'!pre))) @ [item_rule_head (item (bs!k!red))]" + using IH_r by simp + also have "... = take (item_dot (item e)) (item_rule_body (item e))" + using transform by (auto simp: next_symbol_def is_complete_def split: if_splits; metis leI take_Suc_conv_app_nth) + finally have roots: "map root_tree (ts0 @ [Branch N_red ts]) = take (item_dot (item e)) (item_rule_body (item e))" . + have "wf_item \ \ (item (bs!k!red))" + using prems(5,6) bounds(3) unfolding wf_bins_def wf_bin_def wf_bin_items_def by (auto simp: items_def) + moreover have "N_red = item_rule_head (item (bs!k!red))" + using IH_r by fastforce + moreover have "map root_tree ts = item_rule_body (item (bs!k!red))" + using IH_r transform(5) by (auto simp: is_complete_def) + ultimately have "\r \ set (\ \). N_red = rule_head r \ map root_tree ts = rule_body r" + unfolding wf_item_def item_rule_body_def item_rule_head_def by blast + hence "wf_rule_tree \ (Branch N_red ts)" + using IH_r by simp + hence "wf_item_tree \ (item (bs!k!i)) (Branch N (ts0 @ [Branch N_red ts]))" + using transform(1) roots IH_pre entry by simp + thus "wf_item_tree \ (item (bs!k!i)) t" + using tsx(1) red ts by blast + qed + moreover have "?g g = Some fs" + using fs pres rss g by (auto, metis bind.bind_lunit) + ultimately show "\fs. ?g g = Some fs \ (\f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t)" + by blast + qed + ultimately show "\fs. fso = Some fs \ (\f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t)" + using map_Some_P by auto + qed + then obtain fss where "those (map ?g gs) = Some fss" "\fs \ set fss. \f \ set fs. \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t" + using those_Some_P by blast + hence "build_trees' bs \ k i I = Some (concat fss)" "\f \ set (concat fss). \t \ set (trees f). wf_item_tree \ (item (bs!k!i)) t" + using simp by auto + thus ?thesis + using prems(8-10) by auto + qed + qed + done + thus ?thesis + by blast +qed + +lemma wf_yield_tree_build_trees': + assumes "(bs, \, k, i, I) \ wf_trees_input" + assumes "wf_bins \ \ bs" + assumes "k < length bs" "i < length (bs!k)" "k \ length \" + assumes "build_trees' bs \ k i I = Some fs" + assumes "f \ set fs" + assumes "t \ set (trees f)" + shows "wf_yield_tree \ (item (bs!k!i)) t" +proof - + have "wf_yield_tree \ (item (bs!k!i)) t" + using assms + apply (induction arbitrary: fs f t rule: build_trees'_induct[OF assms(1)]) + subgoal premises prems for bs \ k i I fs f t + proof - + define e where entry: "e = bs!k!i" + consider (Null) "pointer e = Null" + | (Pre) "\pre. pointer e = Pre pre" + | (PreRed) "\p ps. pointer e = PreRed p ps" + by (metis pointer.exhaust) + thus ?thesis + proof cases + case Null + hence simp: "build_trees' bs \ k i I = Some ([FBranch (item_rule_head (item e)) []])" + using entry by simp + moreover have "f = FBranch (item_rule_head (item e)) []" + using build_forest'_simps(1) Null prems(9,10) entry by auto + ultimately have simp: "t = Branch (item_rule_head (item e)) []" + using prems(11) by simp + have "sound_ptrs \ bs" + using prems(4) unfolding wf_trees_input_def by blast + hence "predicts (item e)" + using Null prems(6,7) nth_mem entry unfolding sound_ptrs_def sound_null_ptr_def by blast + thus ?thesis + unfolding wf_yield_tree_def predicts_def using simp entry by (auto simp: slice_empty) + next + case Pre + then obtain pre where pre: "pointer e = Pre pre" + by blast + have sound: "sound_ptrs \ bs" + using prems(4) unfolding wf_trees_input_def by blast + hence bounds: "k > 0" "pre < length (bs!(k-1))" + using entry pre prems(6,7) unfolding sound_ptrs_def sound_pre_ptr_def by (metis nth_mem)+ + have scans: "scans \ k (item (bs!(k-1)!pre)) (item e)" + using entry pre prems(6-7) \sound_ptrs \ bs\ unfolding sound_ptrs_def sound_pre_ptr_def by simp + have wf: "(bs, \, k-1, pre, {pre}) \ wf_trees_input" + using entry pre prems(4) wf_trees_input_pre by blast + then obtain pres where pres: "build_trees' bs \ (k-1) pre {pre} = Some pres" + "\f \ set pres. \N fss. f = FBranch N fss" + using build_trees'_termination wf by blast + let ?g = "\f. case f of FBranch N fss \ Some (FBranch N (fss @ [[FLeaf (\!(k-1))]])) | _ \ None" + have "build_trees' bs \ k i I = those (map ?g pres)" + using entry pre pres by simp + hence fs: "Some fs = those (map ?g pres)" + using prems(9) by simp + then obtain f_pre N fss where Nfss: "f = FBranch N (fss @ [[FLeaf (\!(k-1))]])" + "f_pre = FBranch N fss" "f_pre \ set pres" + using those_map_FBranch_only fs pres(2) prems(10) by blast + define tss where tss: "tss = map (\fs. concat (map (\f. trees f) fs)) fss" + have "trees (FBranch N (fss @ [[FLeaf (\!(k-1))]])) = + map (\ts. Branch N ts) [ ts @ [Leaf (\!(k-1))] . ts <- combinations tss ]" + by (subst tss, subst trees_append_single_singleton, simp) + moreover have "t \ set (trees (FBranch N (fss @ [[FLeaf (\!(k-1))]])))" + using Nfss(1) prems(11) by blast + ultimately obtain ts where ts: "t = Branch N (ts @ [Leaf (\!(k-1))]) \ ts \ set (combinations tss)" + by auto + have "sound_ptrs \ bs" + using prems(4) unfolding wf_trees_input_def by blast + hence "pre < length (bs!(k-1))" + using entry pre prems(6,7) unfolding sound_ptrs_def sound_pre_ptr_def by (metis nth_mem) + moreover have "k-1 < length bs" + by (simp add: prems(6) less_imp_diff_less) + moreover have "Branch N ts \ set (trees (FBranch N fss))" + using ts tss by simp + ultimately have IH: "wf_yield_tree \ (item (bs!(k-1)!pre)) (Branch N ts)" + using prems(1,2,4,5,8) entry pre Nfss(2,3) wf pres(1) by simp + have transform: + "item_origin (item (bs!(k-1)!pre)) \ item_end (item (bs!(k-1)!pre))" + "item_end (item (bs!(k-1)!pre)) = k-1" + "item_end (item e) = k" + using entry prems(5,6,7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def wf_item_def + by (auto, meson less_imp_diff_less nth_mem) + have "yield_tree t = concat (map yield_tree (ts @ [Leaf (\!(k-1))]))" + by (simp add: ts) + also have "... = concat (map yield_tree ts) @ [\!(k-1)]" + by simp + also have "... = slice (item_origin (item (bs!(k-1)!pre))) (item_end (item (bs!(k-1)!pre))) \ @ [\!(k-1)]" + using IH by (simp add: wf_yield_tree_def) + also have "... = slice (item_origin (item (bs!(k-1)!pre))) (item_end (item (bs!(k-1)!pre)) + 1) \" + using slice_append_nth transform \k > 0\ prems(8) + by (metis diff_less le_eq_less_or_eq less_imp_diff_less less_numeral_extra(1)) + also have "... = slice (item_origin (item e)) (item_end (item (bs!(k-1)!pre)) + 1) \" + using scans unfolding scans_def inc_item_def by simp + also have "... = slice (item_origin (item e)) k \" + using scans transform unfolding scans_def by (metis Suc_diff_1 Suc_eq_plus1 bounds(1)) + also have "... = slice (item_origin (item e)) (item_end (item e)) \" + using transform by auto + finally show ?thesis + using wf_yield_tree_def entry by blast + next + case PreRed + then obtain p ps where prered: "pointer e = PreRed p ps" + by blast + define ps' where ps': "ps' = filter (\(k', pre, red). red \ I) (p#ps)" + define gs where gs: "gs = group_by (\(k', pre, red). (k', pre)) (\(k', pre, red). red) ps'" + let ?g = "\((k', pre), reds). + do { + pres \ build_trees' bs \ k' pre {pre}; + rss \ those (map (\red. build_trees' bs \ k red (I \ {red})) reds); + those (map (\f. + case f of + FBranch N fss \ Some (FBranch N (fss @ [concat rss])) + | _ \ None \\impossible case\ + ) pres) + }" + have simp: "build_trees' bs \ k i I = map_option concat (those (map ?g gs))" + using entry prered ps' gs by (subst build_trees'.simps) (auto simp del: filter.simps) + have "\fso \ set (map ?g gs). \fs. fso = Some fs \ (\f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t)" + proof standard + fix fso + assume "fso \ set (map ?g gs)" + moreover have "\ps \ set gs. \fs. ?g ps = Some fs \ (\f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t)" + proof standard + fix g + assume "g \ set gs" + then obtain k' pre reds where g: "((k', pre), reds) \ set gs" "((k', pre), reds) = g" + by (metis surj_pair) + moreover have wf_pre: "(bs, \, k', pre, {pre}) \ wf_trees_input" + using wf_trees_input_prered_pre[OF prems(4) entry prered ps' gs g(1)] by blast + ultimately obtain pres where pres: "build_trees' bs \ k' pre {pre} = Some pres" + "\f_pre \ set pres. \N fss. f_pre = FBranch N fss" + using build_trees'_termination by blast + have wf_reds: "\red \ set reds. (bs, \, k, red, I \ {red}) \ wf_trees_input" + using wf_trees_input_prered_red[OF prems(4) entry prered ps' gs g(1)] by blast + hence "\f \ set (map (\red. build_trees' bs \ k red (I \ {red})) reds). \a. f = Some a" + using build_trees'_termination by fastforce + then obtain rss where rss: "Some rss = those (map (\red. build_trees' bs \ k red (I \ {red})) reds)" + using those_Some by (metis (full_types)) + let ?h = "\f. case f of FBranch N fss \ Some (FBranch N (fss @ [concat rss])) | _ \ None" + have "\x \ set (map ?h pres). \a. x = Some a" + using pres(2) by auto + then obtain fs where fs: "Some fs = those (map ?h pres)" + using those_Some by (smt (verit, best)) + have "\f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t" + proof (standard, standard) + fix f t + assume ft: "f \ set fs" "t \ set (trees f)" + hence "\x. x \ set pres \ Some f \ set (map ?h pres)" + using those_map_exists[OF fs ft(1)] by blast + then obtain f_pre N fss where f_pre: "f_pre \ set pres" "f_pre = FBranch N fss" + "f = FBranch N (fss @ [concat rss])" + using pres(2) by force + define tss where tss: "tss = map (\fs. concat (map (\f. trees f) fs)) fss" + have "trees (FBranch N (fss @ [concat rss])) = + map (\ts. Branch N ts) [ ts0 @ ts1 . ts0 <- combinations tss, + ts1 <- combinations [concat (map (\f. trees f) (concat rss)) ] ]" + by (subst tss, subst trees_append_singleton, simp) + moreover have "t \ set (trees (FBranch N (fss @ [concat rss])))" + using ft(2) f_pre(3) by blast + ultimately obtain ts0 ts1 where tsx: "t = Branch N (ts0 @ [ts1])" "ts0 \ set (combinations tss)" + "ts1 \ set (concat (map (\f. trees f) (concat rss)))" + by fastforce + then obtain f_red where f_red: "f_red \ set (concat rss)" "ts1 \ set (trees f_red)" + by auto + obtain fs_red red where red: "Some fs_red = build_trees' bs \ k red (I \ {red})" + "f_red \ set fs_red" "red \ set reds" + using f_red(1) rss those_map_Some_concat_exists by fast + then obtain N_red fss_red where "f_red = FBranch N_red fss_red" + using build_trees'_termination wf_reds by (metis option.inject) + then obtain ts where ts: "Branch N_red ts = ts1" + using tsx(3) f_red by auto + have "(k', pre, red) \ set ps'" + using group_by_forall_v_exists_k \((k', pre), reds) \ set gs\ gs \red \ set reds\ by fast + hence mem: "(k', pre, red) \ set (p#ps)" + using ps' by (metis filter_set member_filter) + have "sound_ptrs \ bs" + using prems(4) wf_trees_input_def by fastforce + have bounds: "k' < k" "pre < length (bs!k')" "red < length (bs!k)" + using prered entry prems(6,7) \sound_ptrs \ bs\ + unfolding sound_ptrs_def sound_prered_ptr_def by (meson mem nth_mem)+ + have completes: "completes k (item (bs!k'!pre)) (item e) (item (bs!k!red))" + using prered entry prems(6,7) \sound_ptrs \ bs\ + unfolding sound_ptrs_def sound_prered_ptr_def by (metis mem nth_mem) + have transform: + "item_rule_head (item (bs!k'!pre)) = item_rule_head (item e)" + "item_rule_body (item (bs!k'!pre)) = item_rule_body (item e)" + "item_dot (item (bs!k'!pre)) + 1 = item_dot (item e)" + "next_symbol (item (bs!k'!pre)) = Some (item_rule_head (item (bs!k!red)))" + "is_complete (item (bs!k!red))" + using completes unfolding completes_def inc_item_def + by (auto simp: item_rule_head_def item_rule_body_def is_complete_def) + have "Branch N ts0 \ set (trees (FBranch N fss))" + using tss tsx(2) by simp + hence IH_pre: "wf_yield_tree \ (item (bs!k'!pre)) (Branch N ts0)" + using prems(2)[OF entry prered ps' gs \((k', pre), reds) \ set gs\ wf_pre prems(5)] + pres(1) f_pre f_pre(3) bounds(1,2) prems(6,8) by simp + have IH_r: "wf_yield_tree \ (item (bs!k!red)) (Branch N_red ts)" + using prems(3)[OF entry prered ps' gs \((k', pre), reds) \ set gs\ \red \ set reds\ _ prems(5)] + bounds(3) f_red(2) red ts wf_reds prems(6,8) by metis + have wf1: + "item_origin (item (bs!k'!pre)) \ item_end (item (bs!k'!pre))" + "item_origin (item (bs!k!red)) \ item_end (item (bs!k!red))" + using prems(5-7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def wf_item_def + by (metis length_map nth_map nth_mem order_less_trans)+ + have wf2: + "item_end (item (bs!k!red)) = k" + "item_end (item (bs!k!i)) = k" + using prems(5-7) bounds unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def by simp_all + have "yield_tree t = concat (map yield_tree (ts0 @ [Branch N_red ts]))" + by (simp add: ts tsx(1)) + also have "... = concat (map yield_tree ts0) @ yield_tree (Branch N_red ts)" + by simp + also have "... = slice (item_origin (item (bs!k'!pre))) (item_end (item (bs!k'!pre))) \ @ + slice (item_origin (item (bs!k!red))) (item_end (item (bs!k!red))) \" + using IH_pre IH_r by (simp add: wf_yield_tree_def) + also have "... = slice (item_origin (item (bs!k'!pre))) (item_end (item (bs!k!red))) \" + using slice_concat wf1 completes_def completes by (metis (no_types, lifting)) + also have "... = slice (item_origin (item e)) (item_end (item (bs!k!red))) \" + using completes unfolding completes_def inc_item_def by simp + also have "... = slice (item_origin (item e)) (item_end (item e)) \" + using wf2 entry by presburger + finally show "wf_yield_tree \ (item (bs!k!i)) t" + using wf_yield_tree_def entry by blast + qed + moreover have "?g g = Some fs" + using fs pres rss g by (auto, metis bind.bind_lunit) + ultimately show "\fs. ?g g = Some fs \ (\f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t)" + by blast + qed + ultimately show "\fs. fso = Some fs \ (\f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t)" + using map_Some_P by auto + qed + then obtain fss where "those (map ?g gs) = Some fss" "\fs \ set fss. \f \ set fs. \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t" + using those_Some_P by blast + hence "build_trees' bs \ k i I = Some (concat fss)" "\f \ set (concat fss). \t \ set (trees f). wf_yield_tree \ (item (bs!k!i)) t" + using simp by auto + thus ?thesis + using prems(9-11) by auto + qed + qed + done + thus ?thesis + using assms(2) by blast +qed + +theorem wf_rule_root_yield_tree_build_trees: + assumes "wf_bins \ \ bs" "sound_ptrs \ bs" "length bs = length \ + 1" + assumes "build_trees \ \ bs = Some fs" "f \ set fs" "t \ set (trees f)" + shows "wf_rule_tree \ t \ root_tree t = \ \ \ yield_tree t = \" +proof - + let ?k = "length bs - 1" + define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items (bs!?k))" + have #: "Some fs = map_option concat (those (map (\(_, i). build_trees' bs \ ?k i {i}) finished))" + using assms(4) build_trees_def finished_def by (metis (full_types)) + then obtain fss fs' where fss: "Some fss = those (map (\(_, i). build_trees' bs \ ?k i {i}) finished)" + "fs' \ set fss" "f \ set fs'" + using map_option_concat_those_map_exists assms(5) by fastforce + then obtain x i where *: "(x,i) \ set finished" "Some fs' = build_trees' bs \ (length bs - 1) i {i}" + using those_map_exists[OF fss(1,2)] by auto + have k: "?k < length bs" "?k \ length \" + using assms(3) by simp_all + have i: "i < length (bs!?k)" + using index_filter_with_index_lt_length * items_def finished_def by (metis (no_types, opaque_lifting) length_map) + have x: "x = item (bs!?k!i)" + using * i filter_with_index_nth items_def nth_map finished_def assms(3) by metis + have finished: "is_finished \ \ x" + using * filter_with_index_P finished_def by metis + have "{i} \ {0.., ?k, i, {i}) \ wf_trees_input" + unfolding wf_trees_input_def using assms(2) i k(1) by simp + hence wf_item_tree: "wf_item_tree \ (item (bs!?k!i)) t" + using wf_item_tree_build_trees' assms(1,2,5,6) i k(1) x *(2) fss(3) by metis + have wf_item: "wf_item \ \ (item (bs!?k!i))" + using k(1) i assms(1) unfolding wf_bins_def wf_bin_def wf_bin_items_def by (simp add: items_def) + obtain N fss where Nfss: "f = FBranch N fss" + using build_trees'_termination[OF wf] by (metis "*"(2) fss(3) option.inject) + then obtain ts where ts: "t = Branch N ts" + using assms(6) by auto + hence "N = item_rule_head x" + "map root_tree ts = item_rule_body x" + using finished wf_item_tree x by (auto simp: is_finished_def is_complete_def) + hence "\r \ set (\ \). N = rule_head r \ map root_tree ts = rule_body r" + using wf_item x unfolding wf_item_def item_rule_body_def item_rule_head_def by blast + hence wf_rule: "wf_rule_tree \ t" + using wf_item_tree ts by simp + have root: "root_tree t = \ \" + using finished ts \N = item_rule_head x\ by (auto simp: is_finished_def) + have "yield_tree t = slice (item_origin (item (bs!?k!i))) (item_end (item (bs!?k!i))) \" + using k i assms(1,6) wf wf_yield_tree_build_trees' wf_yield_tree_def *(2) fss(3) by (smt (verit, best)) + hence yield: "yield_tree t = \" + using finished x unfolding is_finished_def by simp + show ?thesis + using * wf_rule root yield assms(4) unfolding build_trees_def by simp +qed + +corollary wf_rule_root_yield_tree_build_trees_Earley\<^sub>L: + assumes "wf_\ \" "nonempty_derives \" + assumes "build_trees \ \ (Earley\<^sub>L \ \) = Some fs" "f \ set fs" "t \ set (trees f)" + shows "wf_rule_tree \ t \ root_tree t = \ \ \ yield_tree t = \" + using assms wf_rule_root_yield_tree_build_trees wf_bins_Earley\<^sub>L Earley\<^sub>L_def + length_Earley\<^sub>L_bins length_bins_Init\<^sub>L sound_mono_ptrs_Earley\<^sub>L + by (metis dual_order.eq_iff ) + +theorem soundness_build_trees_Earley\<^sub>L: + assumes "wf_\ \" "is_word \ \" "nonempty_derives \" + assumes "build_trees \ \ (Earley\<^sub>L \ \) = Some fs" "f \ set fs" "t \ set (trees f)" + shows "derives \ [\ \] \" +proof - + let ?k = "length (Earley\<^sub>L \ \) - 1" + define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items ((Earley\<^sub>L \ \)!?k))" + have #: "Some fs = map_option concat (those (map (\(_, i). build_trees' (Earley\<^sub>L \ \) \ ?k i {i}) finished))" + using assms(4) build_trees_def finished_def by (metis (full_types)) + then obtain fss fs' where fss: "Some fss = those (map (\(_, i). build_trees' (Earley\<^sub>L \ \) \ ?k i {i}) finished)" + "fs' \ set fss" "f \ set fs'" + using map_option_concat_those_map_exists assms(5) by fastforce + then obtain x i where *: "(x,i) \ set finished" "Some fs' = build_trees' (Earley\<^sub>L \ \) \ ?k i {i}" + using those_map_exists[OF fss(1,2)] by auto + have k: "?k < length (Earley\<^sub>L \ \)" "?k \ length \" + by (simp_all add: Earley\<^sub>L_def assms(1)) + have i: "i < length ((Earley\<^sub>L \ \) ! ?k)" + using index_filter_with_index_lt_length * items_def finished_def by (metis length_map) + have x: "x = item ((Earley\<^sub>L \ \)!?k!i)" + using * i filter_with_index_nth items_def nth_map finished_def by metis + have finished: "is_finished \ \ x" + using * filter_with_index_P finished_def by metis + moreover have "x \ set (items ((Earley\<^sub>L \ \) ! ?k))" + using x by (auto simp: items_def; metis One_nat_def i imageI nth_mem) + ultimately have "recognizing (bins (Earley\<^sub>L \ \)) \ \" + by (meson k(1) kth_bin_sub_bins recognizing_def subsetD) + thus ?thesis + using correctness_Earley\<^sub>L assms by blast +qed + +theorem termination_build_tree_Earley\<^sub>L: + assumes "wf_\ \" "nonempty_derives \" "derives \ [\ \] \" + shows "\fs. build_trees \ \ (Earley\<^sub>L \ \) = Some fs" +proof - + let ?k = "length (Earley\<^sub>L \ \) - 1" + define finished where finished_def: "finished = filter_with_index (is_finished \ \) (items ((Earley\<^sub>L \ \)!?k))" + have "\f \ set finished. (Earley\<^sub>L \ \, \, ?k, snd f, {snd f}) \ wf_trees_input" + proof standard + fix f + assume a: "f \ set finished" + then obtain x i where *: "(x,i) = f" + by (metis surj_pair) + have "sound_ptrs \ (Earley\<^sub>L \ \)" + using sound_mono_ptrs_Earley\<^sub>L assms by blast + moreover have "?k < length (Earley\<^sub>L \ \)" + by (simp add: Earley\<^sub>L_def assms(1)) + moreover have "i < length ((Earley\<^sub>L \ \)!?k)" + using index_filter_with_index_lt_length a * items_def finished_def by (metis length_map) + ultimately show "(Earley\<^sub>L \ \, \, ?k, snd f, {snd f}) \ wf_trees_input" + using * unfolding wf_trees_input_def by auto + qed + hence "\fso \ set (map (\(_, i). build_trees' (Earley\<^sub>L \ \) \ ?k i {i}) finished). \fs. fso = Some fs" + using build_trees'_termination by fastforce + then obtain fss where fss: "Some fss = those (map (\(_, i). build_trees' (Earley\<^sub>L \ \) \ ?k i {i}) finished)" + by (smt (verit, best) those_Some) + then obtain fs where fs: "Some fs = map_option concat (those (map (\(_, i). build_trees' (Earley\<^sub>L \ \) \ ?k i {i}) finished))" + by (metis map_option_eq_Some) + show ?thesis + using finished_def fss fs build_trees_def by (metis (full_types)) +qed + +end \ No newline at end of file diff --git a/thys/Earley_Parser/Earley_Recognizer.thy b/thys/Earley_Parser/Earley_Recognizer.thy new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/Earley_Recognizer.thy @@ -0,0 +1,2331 @@ +theory Earley_Recognizer + imports + Earley_Fixpoint +begin + +section \Earley recognizer\ + +subsection \List auxilaries\ + +fun filter_with_index' :: "nat \ ('a \ bool) \ 'a list \ ('a \ nat) list" where + "filter_with_index' _ _ [] = []" +| "filter_with_index' i P (x#xs) = ( + if P x then (x,i) # filter_with_index' (i+1) P xs + else filter_with_index' (i+1) P xs)" + +definition filter_with_index :: "('a \ bool) \ 'a list \ ('a \ nat) list" where + "filter_with_index P xs = filter_with_index' 0 P xs" + +lemma filter_with_index'_P: + "(x, n) \ set (filter_with_index' i P xs) \ P x" + by (induction xs arbitrary: i) (auto split: if_splits) + +lemma filter_with_index_P: + "(x, n) \ set (filter_with_index P xs) \ P x" + by (metis filter_with_index'_P filter_with_index_def) + +lemma filter_with_index'_cong_filter: + "map fst (filter_with_index' i P xs) = filter P xs" + by (induction xs arbitrary: i) auto + +lemma filter_with_index_cong_filter: + "map fst (filter_with_index P xs) = filter P xs" + by (simp add: filter_with_index'_cong_filter filter_with_index_def) + +lemma size_index_filter_with_index': + "(x, n) \ set (filter_with_index' i P xs) \ n \ i" + by (induction xs arbitrary: i) (auto simp: Suc_leD split: if_splits) + +lemma index_filter_with_index'_lt_length: + "(x, n) \ set (filter_with_index' i P xs) \ n-i < length xs" + by (induction xs arbitrary: i)(auto simp: less_Suc_eq_0_disj split: if_splits; metis Suc_diff_Suc leI)+ + +lemma index_filter_with_index_lt_length: + "(x, n) \ set (filter_with_index P xs) \ n < length xs" + by (metis filter_with_index_def index_filter_with_index'_lt_length minus_nat.diff_0) + +lemma filter_with_index'_nth: + "(x, n) \ set (filter_with_index' i P xs) \ xs ! (n-i) = x" +proof (induction xs arbitrary: i) + case (Cons y xs) + show ?case + proof (cases "x = y") + case True + thus ?thesis + using Cons by (auto simp: nth_Cons' split: if_splits) + next + case False + hence "(x, n) \ set (filter_with_index' (i+1) P xs)" + using Cons.prems by (cases xs) (auto split: if_splits) + hence "n \ i + 1" "xs ! (n - i - 1) = x" + by (auto simp: size_index_filter_with_index' Cons.IH) + thus ?thesis + by simp + qed +qed simp + +lemma filter_with_index_nth: + "(x, n) \ set (filter_with_index P xs) \ xs ! n = x" + by (metis diff_zero filter_with_index'_nth filter_with_index_def) + +lemma filter_with_index_nonempty: + "x \ set xs \ P x \ filter_with_index P xs \ []" + by (metis filter_empty_conv filter_with_index_cong_filter list.map(1)) + +lemma filter_with_index'_Ex_first: + "(\x i xs'. filter_with_index' n P xs = (x, i)#xs') \ (\x \ set xs. P x)" + by (induction xs arbitrary: n) auto + +lemma filter_with_index_Ex_first: + "(\x i xs'. filter_with_index P xs = (x, i)#xs') \ (\x \ set xs. P x)" + using filter_with_index'_Ex_first filter_with_index_def by metis + + +subsection \Definitions\ + +datatype pointer = + Null + | Pre nat \\pre\ + | PreRed "nat \ nat \ nat" "(nat \ nat \ nat) list" \\k', pre, red\ + +datatype 'a entry = + Entry (item : "'a item") (pointer : pointer) + +type_synonym 'a bin = "'a entry list" +type_synonym 'a bins = "'a bin list" + +definition items :: "'a bin \ 'a item list" where + "items b \ map item b" + +definition pointers :: "'a bin \ pointer list" where + "pointers b \ map pointer b" + +definition bins_eq_items :: "'a bins \ 'a bins \ bool" where + "bins_eq_items bs0 bs1 \ map items bs0 = map items bs1" + +definition bins :: "'a bins \ 'a item set" where + "bins bs \ \ { set (items (bs!k)) | k. k < length bs }" + +definition bin_upto :: "'a bin \ nat \ 'a item set" where + "bin_upto b i \ { items b ! j | j. j < i \ j < length (items b) }" + +definition bins_upto :: "'a bins \ nat \ nat \ 'a item set" where + "bins_upto bs k i \ \ { set (items (bs ! l)) | l. l < k } \ bin_upto (bs ! k) i" + +definition wf_bin_items :: "'a cfg \ 'a sentence \ nat \ 'a item list \ bool" where + "wf_bin_items \ \ k xs \ \x \ set xs. wf_item \ \ x \ item_end x = k" + +definition wf_bin :: "'a cfg \ 'a sentence \ nat \ 'a bin \ bool" where + "wf_bin \ \ k b \ distinct (items b) \ wf_bin_items \ \ k (items b)" + +definition wf_bins :: "'a cfg \ 'a list \ 'a bins \ bool" where + "wf_bins \ \ bs \ \k < length bs. wf_bin \ \ k (bs!k)" + +definition nonempty_derives :: "'a cfg \ bool" where + "nonempty_derives \ \ \N. N \ set (\ \) \ \ derives \ [N] []" + +definition Init\<^sub>L :: "'a cfg \ 'a sentence \ 'a bins" where + "Init\<^sub>L \ \ \ + let rs = filter (\r. rule_head r = \ \) (\ \) in + let b0 = map (\r. (Entry (init_item r 0) Null)) rs in + let bs = replicate (length \ + 1) ([]) in + bs[0 := b0]" + +definition Scan\<^sub>L :: "nat \ 'a sentence \ 'a \ 'a item \ nat \ 'a entry list" where + "Scan\<^sub>L k \ a x pre \ + if \!k = a then + let x' = inc_item x (k+1) in + [Entry x' (Pre pre)] + else []" + +definition Predict\<^sub>L :: "nat \ 'a cfg \ 'a \ 'a entry list" where + "Predict\<^sub>L k \ X \ + let rs = filter (\r. rule_head r = X) (\ \) in + map (\r. (Entry (init_item r k) Null)) rs" + +definition Complete\<^sub>L :: "nat \ 'a item \ 'a bins \ nat \ 'a entry list" where + "Complete\<^sub>L k y bs red \ + let orig = bs ! (item_origin y) in + let is = filter_with_index (\x. next_symbol x = Some (item_rule_head y)) (items orig) in + map (\(x, pre). (Entry (inc_item x k) (PreRed (item_origin y, pre, red) []))) is" + +fun bin_upd :: "'a entry \ 'a bin \ 'a bin" where + "bin_upd e' [] = [e']" +| "bin_upd e' (e#es) = ( + case (e', e) of + (Entry x (PreRed px xs), Entry y (PreRed py ys)) \ + if x = y then Entry x (PreRed py (px#xs@ys)) # es + else e # bin_upd e' es + | _ \ + if item e' = item e then e # es + else e # bin_upd e' es)" + +fun bin_upds :: "'a entry list \ 'a bin \ 'a bin" where + "bin_upds [] b = b" +| "bin_upds (e#es) b = bin_upds es (bin_upd e b)" + +definition bins_upd :: "'a bins \ nat \ 'a entry list \ 'a bins" where + "bins_upd bs k es \ bs[k := bin_upds es (bs!k)]" + +partial_function (tailrec) Earley\<^sub>L_bin' :: "nat \ 'a cfg \ 'a sentence \ 'a bins \ nat \ 'a bins" where + "Earley\<^sub>L_bin' k \ \ bs i = ( + if i \ length (items (bs ! k)) then bs + else + let x = items (bs!k) ! i in + let bs' = + case next_symbol x of + Some a \ + if is_terminal \ a then + if k < length \ then bins_upd bs (k+1) (Scan\<^sub>L k \ a x i) + else bs + else bins_upd bs k (Predict\<^sub>L k \ a) + | None \ bins_upd bs k (Complete\<^sub>L k x bs i) + in Earley\<^sub>L_bin' k \ \ bs' (i+1))" + +declare Earley\<^sub>L_bin'.simps[code] + +definition Earley\<^sub>L_bin :: "nat \ 'a cfg \ 'a sentence \ 'a bins \ 'a bins" where + "Earley\<^sub>L_bin k \ \ bs \ Earley\<^sub>L_bin' k \ \ bs 0" + +fun Earley\<^sub>L_bins :: "nat \ 'a cfg \ 'a sentence \ 'a bins" where + "Earley\<^sub>L_bins 0 \ \ = Earley\<^sub>L_bin 0 \ \ (Init\<^sub>L \ \)" +| "Earley\<^sub>L_bins (Suc n) \ \ = Earley\<^sub>L_bin (Suc n) \ \ (Earley\<^sub>L_bins n \ \)" + +definition Earley\<^sub>L :: "'a cfg \ 'a sentence \ 'a bins" where + "Earley\<^sub>L \ \ \ Earley\<^sub>L_bins (length \) \ \" + + +subsection \Bin lemmas\ + +lemma length_bins_upd[simp]: + "length (bins_upd bs k es) = length bs" + unfolding bins_upd_def by simp + +lemma length_bin_upd: + "length (bin_upd e b) \ length b" + by (induction e b rule: bin_upd.induct) (auto split: pointer.splits entry.splits) + +lemma length_bin_upds: + "length (bin_upds es b) \ length b" + by (induction es arbitrary: b) (auto, meson le_trans length_bin_upd) + +lemma length_nth_bin_bins_upd: + "length (bins_upd bs k es ! n) \ length (bs ! n)" + unfolding bins_upd_def using length_bin_upds + by (metis linorder_not_le list_update_beyond nth_list_update_eq nth_list_update_neq order_refl) + +lemma nth_idem_bins_upd: + "k \ n \ bins_upd bs k es ! n = bs ! n" + unfolding bins_upd_def by simp + +lemma items_nth_idem_bin_upd: + "n < length b \ items (bin_upd e b) ! n = items b ! n" + by (induction b arbitrary: e n) (auto simp: items_def less_Suc_eq_0_disj split!: entry.split pointer.split) + +lemma items_nth_idem_bin_upds: + "n < length b \ items (bin_upds es b) ! n = items b ! n" + by (induction es arbitrary: b) + (auto, metis items_def items_nth_idem_bin_upd length_bin_upd nth_map order.strict_trans2) + +lemma items_nth_idem_bins_upd: + "n < length (bs ! k) \ items (bins_upd bs k es ! k) ! n = items (bs ! k) ! n" + unfolding bins_upd_def using items_nth_idem_bin_upds + by (metis linorder_not_less list_update_beyond nth_list_update_eq) + +lemma bin_upto_eq_set_items: + "i \ length b \ bin_upto b i = set (items b)" + by (auto simp: bin_upto_def items_def, metis in_set_conv_nth nth_map order_le_less order_less_trans) + +lemma bins_upto_empty: + "bins_upto bs 0 0 = {}" + unfolding bins_upto_def bin_upto_def by simp + +lemma set_items_bin_upd: + "set (items (bin_upd e b)) = set (items b) \ {item e}" +proof (induction b arbitrary: e) + case (Cons b bs) + show ?case + proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ b = Entry y (PreRed yp ys)") + case True + then obtain x xp xs y yp ys where "e = Entry x (PreRed xp xs)" "b = Entry y (PreRed yp ys)" + by blast + thus ?thesis + using Cons.IH by (auto simp: items_def) + next + case False + then show ?thesis + proof cases + assume *: "item e = item b" + hence "bin_upd e (b # bs) = b # bs" + using False by (auto split: pointer.splits entry.splits) + thus ?thesis + using * by (auto simp: items_def) + next + assume *: "\ item e = item b" + hence "bin_upd e (b # bs) = b # bin_upd e bs" + using False by (auto split: pointer.splits entry.splits) + thus ?thesis + using * Cons.IH by (auto simp: items_def) + qed + qed +qed (auto simp: items_def) + +lemma set_items_bin_upds: + "set (items (bin_upds es b)) = set (items b) \ set (items es)" + using set_items_bin_upd by (induction es arbitrary: b) (auto simp: items_def, blast, force+) + +lemma bins_bins_upd: + assumes "k < length bs" + shows "bins (bins_upd bs k es) = bins bs \ set (items es)" +proof - + let ?bs = "bins_upd bs k es" + have "bins (bins_upd bs k es) = \ {set (items (?bs ! k)) |k. k < length ?bs}" + unfolding bins_def by blast + also have "... = \ {set (items (bs ! l)) |l. l < length bs \ l \ k} \ set (items (?bs ! k))" + unfolding bins_upd_def using assms by (auto, metis nth_list_update) + also have "... = \ {set (items (bs ! l)) |l. l < length bs \ l \ k} \ set (items (bs ! k)) \ set (items es)" + using set_items_bin_upds[of es "bs!k"] by (simp add: assms bins_upd_def sup_assoc) + also have "... = \ {set (items (bs ! k)) |k. k < length bs} \ set (items es)" + using assms by blast + also have "... = bins bs \ set (items es)" + unfolding bins_def by blast + finally show ?thesis . +qed + +lemma kth_bin_sub_bins: + "k < length bs \ set (items (bs ! k)) \ bins bs" + unfolding bins_def bins_upto_def bin_upto_def by blast+ + +lemma bin_upto_Cons_0: + "bin_upto (e#es) 0 = {}" + by (auto simp: bin_upto_def) + +lemma bin_upto_Cons: + assumes "0 < n" + shows "bin_upto (e#es) n = { item e } \ bin_upto es (n-1)" +proof - + have "bin_upto (e#es) n = { items (e#es) ! j | j. j < n \ j < length (items (e#es)) }" + unfolding bin_upto_def by blast + also have "... = { item e } \ { items es ! j | j. j < (n-1) \ j < length (items es) }" + using assms by (cases n) (auto simp: items_def nth_Cons', metis One_nat_def Zero_not_Suc diff_Suc_1 not_less_eq nth_map) + also have "... = { item e } \ bin_upto es (n-1)" + unfolding bin_upto_def by blast + finally show ?thesis . +qed + +lemma bin_upto_nth_idem_bin_upd: + "n < length b \ bin_upto (bin_upd e b) n = bin_upto b n" +proof (induction b arbitrary: e n) + case (Cons b bs) + show ?case + proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ b = Entry y (PreRed yp ys)") + case True + then obtain x xp xs y yp ys where "e = Entry x (PreRed xp xs)" "b = Entry y (PreRed yp ys)" + by blast + thus ?thesis + using Cons bin_upto_Cons_0 + by (cases n) (auto simp: items_def bin_upto_Cons, blast+) + next + case False + then show ?thesis + proof cases + assume *: "item e = item b" + hence "bin_upd e (b # bs) = b # bs" + using False by (auto split: pointer.splits entry.splits) + thus ?thesis + using * by (auto simp: items_def) + next + assume *: "\ item e = item b" + hence "bin_upd e (b # bs) = b # bin_upd e bs" + using False by (auto split: pointer.splits entry.splits) + thus ?thesis + using * Cons bin_upto_Cons_0 + by (cases n) (auto simp: items_def bin_upto_Cons, blast+) + qed + qed +qed (auto simp: items_def) + +lemma bin_upto_nth_idem_bin_upds: + "n < length b \ bin_upto (bin_upds es b) n = bin_upto b n" + using bin_upto_nth_idem_bin_upd length_bin_upd + apply (induction es arbitrary: b) + apply auto + using order.strict_trans2 order.strict_trans1 by blast+ + +lemma bins_upto_kth_nth_idem: + assumes "l < length bs" "k \ l" "n < length (bs ! k)" + shows "bins_upto (bins_upd bs l es) k n = bins_upto bs k n" +proof - + let ?bs = "bins_upd bs l es" + have "bins_upto ?bs k n = \ {set (items (?bs ! l)) |l. l < k} \ bin_upto (?bs ! k) n" + unfolding bins_upto_def by blast + also have "... = \ {set (items (bs ! l)) |l. l < k} \ bin_upto (?bs ! k) n" + unfolding bins_upd_def using assms(1,2) by auto + also have "... = \ {set (items (bs ! l)) |l. l < k} \ bin_upto (bs ! k) n" + unfolding bins_upd_def using assms(1,3) bin_upto_nth_idem_bin_upds + by (metis (no_types, lifting) nth_list_update) + also have "... = bins_upto bs k n" + unfolding bins_upto_def by blast + finally show ?thesis . +qed + +lemma bins_upto_sub_bins: + "k < length bs \ bins_upto bs k n \ bins bs" + unfolding bins_def bins_upto_def bin_upto_def using less_trans by (auto, blast) + +lemma bins_upto_Suc_Un: + "n < length (bs ! k) \ bins_upto bs k (n+1) = bins_upto bs k n \ { items (bs ! k) ! n }" + unfolding bins_upto_def bin_upto_def using less_Suc_eq by (auto simp: items_def, metis nth_map) + +lemma bins_bin_exists: + "x \ bins bs \ \k < length bs. x \ set (items (bs ! k))" + unfolding bins_def by blast + +lemma distinct_bin_upd: + "distinct (items b) \ distinct (items (bin_upd e b))" +proof (induction b arbitrary: e) + case (Cons b bs) + show ?case + proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ b = Entry y (PreRed yp ys)") + case True + then obtain x xp xs y yp ys where "e = Entry x (PreRed xp xs)" "b = Entry y (PreRed yp ys)" + by blast + thus ?thesis + using Cons + apply (auto simp: items_def) + by (metis Un_insert_right entry.sel(1) imageI items_def list.set_map list.simps(15) set_ConsD set_items_bin_upd sup_bot_right) + next + case False + then show ?thesis + proof cases + assume *: "item e = item b" + hence "bin_upd e (b # bs) = b # bs" + using False by (auto split: pointer.splits entry.splits) + thus ?thesis + using * Cons.prems by (auto simp: items_def) + next + assume *: "\ item e = item b" + hence "bin_upd e (b # bs) = b # bin_upd e bs" + using False by (auto split: pointer.splits entry.splits) + moreover have "distinct (items (bin_upd e bs))" + using Cons by (auto simp: items_def) + ultimately show ?thesis + using * Cons.prems set_items_bin_upd + by (metis Un_insert_right distinct.simps(2) insertE items_def list.simps(9) sup_bot_right) + qed + qed +qed (auto simp: items_def) + +lemma wf_bins_kth_bin: + "wf_bins \ \ bs \ k < length bs \ x \ set (items (bs ! k)) \ wf_item \ \ x \ item_end x = k" + using wf_bin_def wf_bins_def wf_bin_items_def by blast + +lemma wf_bin_bin_upd: + assumes "wf_bin \ \ k b" "wf_item \ \ (item e) \ item_end (item e) = k" + shows "wf_bin \ \ k (bin_upd e b)" + using assms +proof (induction b arbitrary: e) + case (Cons b bs) + show ?case + proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ b = Entry y (PreRed yp ys)") + case True + then obtain x xp xs y yp ys where "e = Entry x (PreRed xp xs)" "b = Entry y (PreRed yp ys)" + by blast + thus ?thesis + using Cons distinct_bin_upd wf_bin_def wf_bin_items_def set_items_bin_upd + by (smt (verit, best) Un_insert_right insertE sup_bot.right_neutral) + next + case False + then show ?thesis + proof cases + assume *: "item e = item b" + hence "bin_upd e (b # bs) = b # bs" + using False by (auto split: pointer.splits entry.splits) + thus ?thesis + using * Cons.prems by (auto simp: items_def) + next + assume *: "\ item e = item b" + hence "bin_upd e (b # bs) = b # bin_upd e bs" + using False by (auto split: pointer.splits entry.splits) + thus ?thesis + using * Cons.prems set_items_bin_upd distinct_bin_upd wf_bin_def wf_bin_items_def + by (smt (verit, best) Un_insert_right insertE sup_bot_right) + qed + qed +qed (auto simp: items_def wf_bin_def wf_bin_items_def) + +lemma wf_bin_bin_upds: + assumes "wf_bin \ \ k b" "distinct (items es)" + assumes "\x \ set (items es). wf_item \ \ x \ item_end x = k" + shows "wf_bin \ \ k (bin_upds es b)" + using assms by (induction es arbitrary: b) (auto simp: wf_bin_bin_upd items_def) + +lemma wf_bins_bins_upd: + assumes "wf_bins \ \ bs" "distinct (items es)" + assumes "\x \ set (items es). wf_item \ \ x \ item_end x = k" + shows "wf_bins \ \ (bins_upd bs k es)" + unfolding bins_upd_def using assms wf_bin_bin_upds wf_bins_def + by (metis length_list_update nth_list_update_eq nth_list_update_neq) + +lemma wf_bins_impl_wf_items: + "wf_bins \ \ bs \ \x \ (bins bs). wf_item \ \ x" + unfolding wf_bins_def wf_bin_def wf_bin_items_def bins_def by auto + +lemma bin_upds_eq_items: + "set (items es) \ set (items b) \ set (items (bin_upds es b)) = set (items b)" + apply (induction es arbitrary: b) + apply (auto simp: set_items_bin_upd set_items_bin_upds) + apply (simp add: items_def) + by (metis Un_iff Un_subset_iff items_def list.simps(9) set_subset_Cons) + +lemma bin_eq_items_bin_upd: + "item e \ set (items b) \ items (bin_upd e b) = items b" +proof (induction b arbitrary: e) + case (Cons b bs) + show ?case + proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ b = Entry y (PreRed yp ys)") + case True + then obtain x xp xs y yp ys where "e = Entry x (PreRed xp xs)" "b = Entry y (PreRed yp ys)" + by blast + thus ?thesis + using Cons by (auto simp: items_def) + next + case False + then show ?thesis + proof cases + assume *: "item e = item b" + hence "bin_upd e (b # bs) = b # bs" + using False by (auto split: pointer.splits entry.splits) + thus ?thesis + using * Cons.prems by (auto simp: items_def) + next + assume *: "\ item e = item b" + hence "bin_upd e (b # bs) = b # bin_upd e bs" + using False by (auto split: pointer.splits entry.splits) + thus ?thesis + using * Cons by (auto simp: items_def) + qed + qed +qed (auto simp: items_def) + +lemma bin_eq_items_bin_upds: + assumes "set (items es) \ set (items b)" + shows "items (bin_upds es b) = items b" + using assms +proof (induction es arbitrary: b) + case (Cons e es) + have "items (bin_upds es (bin_upd e b)) = items (bin_upd e b)" + using Cons bin_upds_eq_items set_items_bin_upd set_items_bin_upds + by (metis Un_upper2 bin_upds.simps(2) sup.coboundedI1) + moreover have "items (bin_upd e b) = items b" + using bin_eq_items_bin_upd Cons.prems by (auto simp: items_def) + ultimately show ?case + by simp +qed (auto simp: items_def) + +lemma bins_eq_items_bins_upd: + assumes "set (items es) \ set (items (bs!k))" + shows "bins_eq_items (bins_upd bs k es) bs" + unfolding bins_upd_def using assms bin_eq_items_bin_upds bins_eq_items_def + by (metis list_update_id map_update) + +lemma bins_eq_items_imp_eq_bins: + "bins_eq_items bs bs' \ bins bs = bins bs'" + unfolding bins_eq_items_def bins_def items_def + by (metis (no_types, lifting) length_map nth_map) + +lemma bin_eq_items_dist_bin_upd_bin: + assumes "items a = items b" + shows "items (bin_upd e a) = items (bin_upd e b)" + using assms +proof (induction a arbitrary: e b) + case (Cons a as) + obtain b' bs where bs: "b = b' # bs" "item a = item b'" "items as = items bs" + using Cons.prems by (auto simp: items_def) + show ?case + proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ a = Entry y (PreRed yp ys)") + case True + then obtain x xp xs y yp ys where #: "e = Entry x (PreRed xp xs)" "a = Entry y (PreRed yp ys)" + by blast + show ?thesis + proof cases + assume *: "x = y" + hence "items (bin_upd e (a # as)) = x # items as" + using # by (auto simp: items_def) + moreover have "items (bin_upd e (b' # bs)) = x # items bs" + using bs # * by (auto simp: items_def split: pointer.splits entry.splits) + ultimately show ?thesis + using bs by simp + next + assume *: "\ x = y" + hence "items (bin_upd e (a # as)) = y # items (bin_upd e as)" + using # by (auto simp: items_def) + moreover have "items (bin_upd e (b' # bs)) = y # items (bin_upd e bs)" + using bs # * by (auto simp: items_def split: pointer.splits entry.splits) + ultimately show ?thesis + using bs Cons.IH by simp + qed + next + case False + then show ?thesis + proof cases + assume *: "item e = item a" + hence "items (bin_upd e (a # as)) = item a # items as" + using False by (auto simp: items_def split: pointer.splits entry.splits) + moreover have "items (bin_upd e (b' # bs)) = item b' # items bs" + using bs False * by (auto simp: items_def split: pointer.splits entry.splits) + ultimately show ?thesis + using bs by simp + next + assume *: "\ item e = item a" + hence "items (bin_upd e (a # as)) = item a # items (bin_upd e as)" + using False by (auto simp: items_def split: pointer.splits entry.splits) + moreover have "items (bin_upd e (b' # bs)) = item b' # items (bin_upd e bs)" + using bs False * by (auto simp: items_def split: pointer.splits entry.splits) + ultimately show ?thesis + using bs Cons by simp + qed + qed +qed (auto simp: items_def) + +lemma bin_eq_items_dist_bin_upds_bin: + assumes "items a = items b" + shows "items (bin_upds es a) = items (bin_upds es b)" + using assms +proof (induction es arbitrary: a b) + case (Cons e es) + hence "items (bin_upds es (bin_upd e a)) = items (bin_upds es (bin_upd e b))" + using bin_eq_items_dist_bin_upd_bin by blast + thus ?case + by simp +qed simp + +lemma bin_eq_items_dist_bin_upd_entry: + assumes "item e = item e'" + shows "items (bin_upd e b) = items (bin_upd e' b)" + using assms +proof (induction b arbitrary: e e') + case (Cons a as) + show ?case + proof (cases "\x xp xs y yp ys. e = Entry x (PreRed xp xs) \ a = Entry y (PreRed yp ys)") + case True + then obtain x xp xs y yp ys where #: "e = Entry x (PreRed xp xs)" "a = Entry y (PreRed yp ys)" + by blast + show ?thesis + proof cases + assume *: "x = y" + thus ?thesis + using # Cons.prems by (auto simp: items_def split: pointer.splits entry.splits) + next + assume *: "\ x = y" + thus ?thesis + using # Cons.prems + by (auto simp: items_def split!: pointer.splits entry.splits, metis Cons.IH Cons.prems items_def)+ + qed + next + case False + then show ?thesis + proof cases + assume *: "item e = item a" + thus ?thesis + using Cons.prems by (auto simp: items_def split: pointer.splits entry.splits) + next + assume *: "\ item e = item a" + thus ?thesis + using Cons.prems + by (auto simp: items_def split!: pointer.splits entry.splits, metis Cons.IH Cons.prems items_def)+ + qed + qed +qed (auto simp: items_def) + +lemma bin_eq_items_dist_bin_upds_entries: + assumes "items es = items es'" + shows "items (bin_upds es b) = items (bin_upds es' b)" + using assms +proof (induction es arbitrary: es' b) + case (Cons e es) + then obtain e' es'' where "item e = item e'" "items es = items es''" "es' = e' # es''" + by (auto simp: items_def) + hence "items (bin_upds es (bin_upd e b)) = items (bin_upds es'' (bin_upd e' b))" + using Cons.IH + by (metis bin_eq_items_dist_bin_upd_entry bin_eq_items_dist_bin_upds_bin) + thus ?case + by (simp add: \es' = e' # es''\) +qed (auto simp: items_def) + +lemma bins_eq_items_dist_bins_upd: + assumes "bins_eq_items as bs" "items aes = items bes" "k < length as" + shows "bins_eq_items (bins_upd as k aes) (bins_upd bs k bes)" +proof - + have "k < length bs" + using assms(1,3) bins_eq_items_def map_eq_imp_length_eq by metis + hence "items (bin_upds (as!k) aes) = items (bin_upds (bs!k) bes)" + using bin_eq_items_dist_bin_upds_entries bin_eq_items_dist_bin_upds_bin bins_eq_items_def assms + by (metis (no_types, lifting) nth_map) + thus ?thesis + using \k < length bs\ assms bin_eq_items_dist_bin_upds_bin bin_eq_items_dist_bin_upds_entries + bins_eq_items_def bins_upd_def by (smt (verit) map_update nth_map) +qed + +subsection \Well-formed bins\ + +lemma distinct_Scan\<^sub>L: + "distinct (items (Scan\<^sub>L k \ a x pre))" + unfolding Scan\<^sub>L_def by (auto simp: items_def) + +lemma distinct_Predict\<^sub>L: + "wf_\ \ \ distinct (items (Predict\<^sub>L k \ X))" + unfolding Predict\<^sub>L_def wf_\_defs by (auto simp: init_item_def rule_head_def distinct_map inj_on_def items_def) + +lemma inj_on_inc_item: + "\x \ A. item_end x = l \ inj_on (\x. inc_item x k) A" + unfolding inj_on_def inc_item_def by (simp add: item.expand) + +lemma distinct_Complete\<^sub>L: + assumes "wf_bins \ \ bs" "item_origin y < length bs" + shows "distinct (items (Complete\<^sub>L k y bs red))" +proof - + let ?orig = "bs ! (item_origin y)" + let ?is = "filter_with_index (\x. next_symbol x = Some (item_rule_head y)) (items ?orig)" + let ?is' = "map (\(x, pre). (Entry (inc_item x k) (PreRed (item_origin y, pre, red) []))) ?is" + have wf: "wf_bin \ \ (item_origin y) ?orig" + using assms wf_bins_def by blast + have 0: "\x \ set (map fst ?is). item_end x = (item_origin y)" + using wf wf_bin_def wf_bin_items_def filter_is_subset filter_with_index_cong_filter by (metis in_mono) + hence "distinct (items ?orig)" + using wf unfolding wf_bin_def by blast + hence "distinct (map fst ?is)" + using filter_with_index_cong_filter distinct_filter by metis + moreover have "items ?is' = map (\x. inc_item x k) (map fst ?is)" + by (induction ?is) (auto simp: items_def) + moreover have "inj_on (\x. inc_item x k) (set (map fst ?is))" + using inj_on_inc_item 0 by blast + ultimately have "distinct (items ?is')" + using distinct_map by metis + thus ?thesis + unfolding Complete\<^sub>L_def by simp +qed + +lemma wf_bins_Scan\<^sub>L': + assumes "wf_bins \ \ bs" "k < length bs" "x \ set (items (bs ! k))" + assumes "k < length \" "next_symbol x \ None" "y = inc_item x (k+1)" + shows "wf_item \ \ y \ item_end y = k+1" + using assms wf_bins_kth_bin[OF assms(1-3)] + unfolding wf_item_def inc_item_def next_symbol_def is_complete_def item_rule_body_def + by (auto split: if_splits) + +lemma wf_bins_Scan\<^sub>L: + assumes "wf_bins \ \ bs" "k < length bs" "x \ set (items (bs ! k))" "k < length \" "next_symbol x \ None" + shows "\y \ set (items (Scan\<^sub>L k \ a x pre)). wf_item \ \ y \ item_end y = (k+1)" + using wf_bins_Scan\<^sub>L'[OF assms] by (simp add: Scan\<^sub>L_def items_def) + +lemma wf_bins_Predict\<^sub>L: + assumes "wf_bins \ \ bs" "k < length bs" "k \ length \" "wf_\ \" + shows "\y \ set (items (Predict\<^sub>L k \ X)). wf_item \ \ y \ item_end y = k" + using assms by (auto simp: Predict\<^sub>L_def wf_item_def wf_bins_def wf_bin_def init_item_def wf_\_defs items_def) + +lemma wf_item_inc_item: + assumes "wf_item \ \ x" "next_symbol x = Some a" "item_origin x \ k" "k \ length \" + shows "wf_item \ \ (inc_item x k) \ item_end (inc_item x k) = k" + using assms by (auto simp: wf_item_def inc_item_def item_rule_body_def next_symbol_def is_complete_def split: if_splits) + +lemma wf_bins_Complete\<^sub>L: + assumes "wf_bins \ \ bs" "k < length bs" "y \ set (items (bs ! k))" + shows "\x \ set (items (Complete\<^sub>L k y bs red)). wf_item \ \ x \ item_end x = k" +proof - + let ?orig = "bs ! (item_origin y)" + let ?is = "filter_with_index (\x. next_symbol x = Some (item_rule_head y)) (items ?orig)" + let ?is' = "map (\(x, pre). (Entry (inc_item x k) (PreRed (item_origin y, pre, red) []))) ?is" + { + fix x + assume *: "x \ set (map fst ?is)" + have "item_end x = item_origin y" + using * assms wf_bins_kth_bin wf_item_def filter_with_index_cong_filter + by (metis dual_order.strict_trans2 filter_is_subset subsetD) + have "wf_item \ \ x" + using * assms wf_bins_kth_bin wf_item_def filter_with_index_cong_filter + by (metis dual_order.strict_trans2 filter_is_subset subsetD) + moreover have "next_symbol x = Some (item_rule_head y)" + using * filter_set filter_with_index_cong_filter member_filter by metis + moreover have "item_origin x \ k" + using \item_end x = item_origin y\ \wf_item \ \ x\ assms wf_bins_kth_bin wf_item_def + by (metis dual_order.order_iff_strict dual_order.strict_trans1) + moreover have "k \ length \" + using assms wf_bins_kth_bin wf_item_def by blast + ultimately have "wf_item \ \ (inc_item x k)" "item_end (inc_item x k) = k" + by (simp_all add: wf_item_inc_item) + } + hence "\x \ set (items ?is'). wf_item \ \ x \ item_end x = k" + by (auto simp: items_def rev_image_eqI) + thus ?thesis + unfolding Complete\<^sub>L_def by presburger +qed + +lemma Ex_wf_bins: + "\n bs \ \. n \ length \ \ length bs = Suc (length \) \ wf_\ \ \ wf_bins \ \ bs" + apply (rule exI[where x="0"]) + apply (rule exI[where x="[[]]"]) + apply (rule exI[where x="[]"]) + apply (auto simp: wf_bins_def wf_bin_def wf_\_defs wf_bin_items_def items_def split: prod.splits) + by (metis cfg.sel distinct.simps(1) empty_iff empty_set inf_bot_right list.set_intros(1)) + +definition wf_earley_input :: "(nat \ 'a cfg \ 'a sentence \ 'a bins) set" where + "wf_earley_input = { + (k, \, \, bs) | k \ \ bs. + k \ length \ \ + length bs = length \ + 1 \ + wf_\ \ \ + wf_bins \ \ bs + }" + +typedef 'a wf_bins = "wf_earley_input::(nat \ 'a cfg \ 'a sentence \ 'a bins) set" + morphisms from_wf_bins to_wf_bins + using Ex_wf_bins by (auto simp: wf_earley_input_def) + +lemma wf_earley_input_elim: + assumes "(k, \, \, bs) \ wf_earley_input" + shows "k \ length \ \ k < length bs \ length bs = length \ + 1 \ wf_\ \ \ wf_bins \ \ bs" + using assms(1) from_wf_bins wf_earley_input_def by (smt (verit) Suc_eq_plus1 less_Suc_eq_le mem_Collect_eq prod.sel(1) snd_conv) + +lemma wf_earley_input_intro: + assumes "k \ length \" "length bs = length \ + 1" "wf_\ \" "wf_bins \ \ bs" + shows "(k, \, \, bs) \ wf_earley_input" + by (simp add: assms wf_earley_input_def) + +lemma wf_earley_input_Complete\<^sub>L: + assumes "(k, \, \, bs) \ wf_earley_input" "\ length (items (bs ! k)) \ i" + assumes "x = items (bs ! k) ! i" "next_symbol x = None" + shows "(k, \, \, bins_upd bs k (Complete\<^sub>L k x bs red)) \ wf_earley_input" +proof - + have *: "k \ length \" "length bs = length \ + 1" "wf_\ \" "wf_bins \ \ bs" + using wf_earley_input_elim assms(1) by metis+ + have x: "x \ set (items (bs ! k))" + using assms(2,3) by simp + have "item_origin x < length bs" + using x wf_bins_kth_bin *(1,2,4) wf_item_def + by (metis One_nat_def add.right_neutral add_Suc_right dual_order.trans le_imp_less_Suc) + hence "wf_bins \ \ (bins_upd bs k (Complete\<^sub>L k x bs red))" + using *(1,2,4) Suc_eq_plus1 distinct_Complete\<^sub>L le_imp_less_Suc wf_bins_Complete\<^sub>L wf_bins_bins_upd x by metis + thus ?thesis + by (simp add: *(1-3) wf_earley_input_def) +qed + +lemma wf_earley_input_Scan\<^sub>L: + assumes "(k, \, \, bs) \ wf_earley_input" "\ length (items (bs ! k)) \ i" + assumes "x = items (bs ! k) ! i" "next_symbol x = Some a" + assumes "is_terminal \ a" "k < length \" + shows "(k, \, \, bins_upd bs (k+1) (Scan\<^sub>L k \ a x pre)) \ wf_earley_input" +proof - + have *: "k \ length \" "length bs = length \ + 1" "wf_\ \" "wf_bins \ \ bs" + using wf_earley_input_elim assms(1) by metis+ + have x: "x \ set (items(bs ! k))" + using assms(2,3) by simp + have "wf_bins \ \ (bins_upd bs (k+1) (Scan\<^sub>L k \ a x pre))" + using * x assms(1,4,6) distinct_Scan\<^sub>L wf_bins_Scan\<^sub>L wf_bins_bins_upd wf_earley_input_elim + by (metis option.discI) + thus ?thesis + by (simp add: *(1-3) wf_earley_input_def) +qed + +lemma wf_earley_input_Predict\<^sub>L: + assumes "(k, \, \, bs) \ wf_earley_input" "\ length (items (bs ! k)) \ i" + assumes "x = items (bs ! k) ! i" "next_symbol x = Some a" "\ is_terminal \ a" + shows "(k, \, \, bins_upd bs k (Predict\<^sub>L k \ a)) \ wf_earley_input" +proof - + have *: "k \ length \" "length bs = length \ + 1" "wf_\ \" "wf_bins \ \ bs" + using wf_earley_input_elim assms(1) by metis+ + have x: "x \ set (items (bs ! k))" + using assms(2,3) by simp + hence "wf_bins \ \ (bins_upd bs k (Predict\<^sub>L k \ a))" + using * x assms(1,4) distinct_Predict\<^sub>L wf_bins_Predict\<^sub>L wf_bins_bins_upd wf_earley_input_elim by metis + thus ?thesis + by (simp add: *(1-3) wf_earley_input_def) +qed + +fun earley_measure :: "nat \ 'a cfg \ 'a sentence \ 'a bins \ nat \ nat" where + "earley_measure (k, \, \, bs) i = card { x | x. wf_item \ \ x \ item_end x = k } - i" + +lemma Earley\<^sub>L_bin'_simps[simp]: + "i \ length (items (bs ! k)) \ Earley\<^sub>L_bin' k \ \ bs i = bs" + "\ i \ length (items (bs ! k)) \ x = items (bs!k) ! i \ next_symbol x = None \ + Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ (bins_upd bs k (Complete\<^sub>L k x bs i)) (i+1)" + "\ i \ length (items (bs ! k)) \ x = items (bs!k) ! i \ next_symbol x = Some a \ + is_terminal \ a \ k < length \ \ Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ (bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)) (i+1)" + "\ i \ length (items (bs ! k)) \ x = items (bs!k) ! i \ next_symbol x = Some a \ + is_terminal \ a \ \ k < length \ \ Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ bs (i+1)" + "\ i \ length (items (bs ! k)) \ x = items (bs!k) ! i \ next_symbol x = Some a \ + \ is_terminal \ a \ Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ (bins_upd bs k (Predict\<^sub>L k \ a)) (i+1)" + by (subst Earley\<^sub>L_bin'.simps, simp)+ + +lemma Earley\<^sub>L_bin'_induct[case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes base: "\k \ \ bs i. i \ length (items (bs ! k)) \ P k \ \ bs i" + assumes complete: "\k \ \ bs i x. \ i \ length (items (bs ! k)) \ x = items (bs ! k) ! i \ + next_symbol x = None \ P k \ \ (bins_upd bs k (Complete\<^sub>L k x bs i)) (i+1) \ P k \ \ bs i" + assumes scan: "\k \ \ bs i x a. \ i \ length (items (bs ! k)) \ x = items (bs ! k) ! i \ + next_symbol x = Some a \ is_terminal \ a \ k < length \ \ + P k \ \ (bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)) (i+1) \ P k \ \ bs i" + assumes pass: "\k \ \ bs i x a. \ i \ length (items (bs ! k)) \ x = items (bs ! k) ! i \ + next_symbol x = Some a \ is_terminal \ a \ \ k < length \ \ + P k \ \ bs (i+1) \ P k \ \ bs i" + assumes predict: "\k \ \ bs i x a. \ i \ length (items (bs ! k)) \ x = items (bs ! k) ! i \ + next_symbol x = Some a \ \ is_terminal \ a \ + P k \ \ (bins_upd bs k (Predict\<^sub>L k \ a)) (i+1) \ P k \ \ bs i" + shows "P k \ \ bs i" + using assms(1) +proof (induction n\"earley_measure (k, \, \, bs) i" arbitrary: bs i rule: nat_less_induct) + case 1 + have wf: "k \ length \" "length bs = length \ + 1" "wf_\ \" "wf_bins \ \ bs" + using "1.prems" wf_earley_input_elim by metis+ + hence k: "k < length bs" + by simp + have fin: "finite { x | x. wf_item \ \ x \ item_end x = k }" + using finiteness_UNIV_wf_item by fastforce + show ?case + proof cases + assume "i \ length (items (bs ! k))" + then show ?thesis + by (simp add: base) + next + assume a1: "\ i \ length (items (bs ! k))" + let ?x = "items (bs ! k) ! i" + have x: "?x \ set (items (bs ! k))" + using a1 by fastforce + show ?thesis + proof cases + assume a2: "next_symbol ?x = None" + let ?bs' = "bins_upd bs k (Complete\<^sub>L k ?x bs i)" + have "item_origin ?x < length bs" + using wf(4) k wf_bins_kth_bin wf_item_def x by (metis order_le_less_trans) + hence wf_bins': "wf_bins \ \ ?bs'" + using wf_bins_Complete\<^sub>L distinct_Complete\<^sub>L wf(4) wf_bins_bins_upd k x by metis + hence wf': "(k, \, \, ?bs') \ wf_earley_input" + using wf(1,2,3) wf_earley_input_intro by fastforce + have sub: "set (items (?bs' ! k)) \ { x | x. wf_item \ \ x \ item_end x = k }" + using wf(1,2) wf_bins' unfolding wf_bin_def wf_bins_def wf_bin_items_def using order_le_less_trans by auto + have "i < length (items (?bs' ! k))" + using a1 by (metis dual_order.strict_trans1 items_def leI length_map length_nth_bin_bins_upd) + also have "... = card (set (items (?bs' ! k)))" + using wf(1,2) wf_bins' distinct_card wf_bins_def wf_bin_def by (metis k length_bins_upd) + also have "... \ card {x |x. wf_item \ \ x \ item_end x = k}" + using card_mono fin sub by blast + finally have "card {x |x. wf_item \ \ x \ item_end x = k} > i" + by blast + hence "earley_measure (k, \, \, ?bs') (Suc i) < earley_measure (k, \, \, bs) i" + by simp + thus ?thesis + using 1 a1 a2 complete wf' by simp + next + assume a2: "\ next_symbol ?x = None" + then obtain a where a_def: "next_symbol ?x = Some a" + by blast + show ?thesis + proof cases + assume a3: "is_terminal \ a" + show ?thesis + proof cases + assume a4: "k < length \" + let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a ?x i)" + have wf_bins': "wf_bins \ \ ?bs'" + using wf_bins_Scan\<^sub>L distinct_Scan\<^sub>L wf(1,4) wf_bins_bins_upd a2 a4 k x by metis + hence wf': "(k, \, \, ?bs') \ wf_earley_input" + using wf(1,2,3) wf_earley_input_intro by fastforce + have sub: "set (items (?bs' ! k)) \ { x | x. wf_item \ \ x \ item_end x = k }" + using wf(1,2) wf_bins' unfolding wf_bin_def wf_bins_def wf_bin_items_def using order_le_less_trans by auto + have "i < length (items (?bs' ! k))" + using a1 by (metis dual_order.strict_trans1 items_def leI length_map length_nth_bin_bins_upd) + also have "... = card (set (items (?bs' ! k)))" + using wf(1,2) wf_bins' distinct_card wf_bins_def wf_bin_def + by (metis Suc_eq_plus1 le_imp_less_Suc length_bins_upd) + also have "... \ card {x |x. wf_item \ \ x \ item_end x = k}" + using card_mono fin sub by blast + finally have "card {x |x. wf_item \ \ x \ item_end x = k} > i" + by blast + hence "earley_measure (k, \, \, ?bs') (Suc i) < earley_measure (k, \, \, bs) i" + by simp + thus ?thesis + using 1 a1 a_def a3 a4 scan wf' by simp + next + assume a4: "\ k < length \" + have sub: "set (items (bs ! k)) \ { x | x. wf_item \ \ x \ item_end x = k }" + using wf(1,2,4) unfolding wf_bin_def wf_bins_def wf_bin_items_def using order_le_less_trans by auto + have "i < length (items (bs ! k))" + using a1 by simp + also have "... = card (set (items (bs ! k)))" + using wf(1,2,4) distinct_card wf_bins_def wf_bin_def by (metis Suc_eq_plus1 le_imp_less_Suc) + also have "... \ card {x |x. wf_item \ \ x \ item_end x = k}" + using card_mono fin sub by blast + finally have "card {x |x. wf_item \ \ x \ item_end x = k} > i" + by blast + hence "earley_measure (k, \, \, bs) (Suc i) < earley_measure (k, \, \, bs) i" + by simp + thus ?thesis + using 1 a1 a3 a4 a_def pass by simp + qed + next + assume a3: "\ is_terminal \ a" + let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" + have wf_bins': "wf_bins \ \ ?bs'" + using wf_bins_Predict\<^sub>L distinct_Predict\<^sub>L wf(1,3,4) wf_bins_bins_upd k x by metis + hence wf': "(k, \, \, ?bs') \ wf_earley_input" + using wf(1,2,3) wf_earley_input_intro by fastforce + have sub: "set (items (?bs' ! k)) \ { x | x. wf_item \ \ x \ item_end x = k }" + using wf(1,2) wf_bins' unfolding wf_bin_def wf_bins_def wf_bin_items_def using order_le_less_trans by auto + have "i < length (items (?bs' ! k))" + using a1 by (metis dual_order.strict_trans1 items_def leI length_map length_nth_bin_bins_upd) + also have "... = card (set (items (?bs' ! k)))" + using wf(1,2) wf_bins' distinct_card wf_bins_def wf_bin_def + by (metis Suc_eq_plus1 le_imp_less_Suc length_bins_upd) + also have "... \ card {x |x. wf_item \ \ x \ item_end x = k}" + using card_mono fin sub by blast + finally have "card {x |x. wf_item \ \ x \ item_end x = k} > i" + by blast + hence "earley_measure (k, \, \, ?bs') (Suc i) < earley_measure (k, \, \, bs) i" + by simp + thus ?thesis + using 1 a1 a_def a3 a_def predict wf' by simp + qed + qed + qed +qed + +lemma wf_earley_input_Earley\<^sub>L_bin': + assumes "(k, \, \, bs) \ wf_earley_input" + shows "(k, \, \, Earley\<^sub>L_bin' k \ \ bs i) \ wf_earley_input" + using assms +proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) + case (Complete\<^sub>F k \ \ bs i x) + let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" + have "(k, \, \, ?bs') \ wf_earley_input" + using Complete\<^sub>F.hyps Complete\<^sub>F.prems wf_earley_input_Complete\<^sub>L by blast + thus ?case + using Complete\<^sub>F.IH Complete\<^sub>F.hyps by simp +next + case (Scan\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" + have "(k, \, \, ?bs') \ wf_earley_input" + using Scan\<^sub>F.hyps Scan\<^sub>F.prems wf_earley_input_Scan\<^sub>L by metis + thus ?case + using Scan\<^sub>F.IH Scan\<^sub>F.hyps by simp +next + case (Predict\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" + have "(k, \, \, ?bs') \ wf_earley_input" + using Predict\<^sub>F.hyps Predict\<^sub>F.prems wf_earley_input_Predict\<^sub>L by metis + thus ?case + using Predict\<^sub>F.IH Predict\<^sub>F.hyps by simp +qed simp_all + +lemma wf_earley_input_Earley\<^sub>L_bin: + assumes "(k, \, \, bs) \ wf_earley_input" + shows "(k, \, \, Earley\<^sub>L_bin k \ \ bs) \ wf_earley_input" + using assms by (simp add: Earley\<^sub>L_bin_def wf_earley_input_Earley\<^sub>L_bin') + +lemma length_bins_Earley\<^sub>L_bin': + assumes "(k, \, \, bs) \ wf_earley_input" + shows "length (Earley\<^sub>L_bin' k \ \ bs i) = length bs" + by (metis assms wf_earley_input_Earley\<^sub>L_bin' wf_earley_input_elim) + +lemma length_nth_bin_Earley\<^sub>L_bin': + assumes "(k, \, \, bs) \ wf_earley_input" + shows "length (items (Earley\<^sub>L_bin' k \ \ bs i ! l)) \ length (items (bs ! l))" + using length_nth_bin_bins_upd order_trans + by (induction i rule: Earley\<^sub>L_bin'_induct[OF assms]) (auto simp: items_def, blast+) + +lemma wf_bins_Earley\<^sub>L_bin': + assumes "(k, \, \, bs) \ wf_earley_input" + shows "wf_bins \ \ (Earley\<^sub>L_bin' k \ \ bs i)" + using assms wf_earley_input_Earley\<^sub>L_bin' wf_earley_input_elim by blast + +lemma wf_bins_Earley\<^sub>L_bin: + assumes "(k, \, \, bs) \ wf_earley_input" + shows "wf_bins \ \ (Earley\<^sub>L_bin k \ \ bs)" + using assms Earley\<^sub>L_bin_def wf_bins_Earley\<^sub>L_bin' by metis + +lemma kth_Earley\<^sub>L_bin'_bins: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "j < length (items (bs ! l))" + shows "items (Earley\<^sub>L_bin' k \ \ bs i ! l) ! j = items (bs ! l) ! j" + using assms(2) +proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) + case (Complete\<^sub>F k \ \ bs i x) + let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" + have "items (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1) ! l) ! j = items (?bs' ! l) ! j" + using Complete\<^sub>F.IH Complete\<^sub>F.prems length_nth_bin_bins_upd items_def order.strict_trans2 by (metis length_map) + also have "... = items (bs ! l) ! j" + using Complete\<^sub>F.prems items_nth_idem_bins_upd nth_idem_bins_upd length_map items_def by metis + finally show ?case + using Complete\<^sub>F.hyps by simp +next + case (Scan\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" + have "items (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1) ! l) ! j = items (?bs' ! l) ! j" + using Scan\<^sub>F.IH Scan\<^sub>F.prems length_nth_bin_bins_upd order.strict_trans2 items_def by (metis length_map) + also have "... = items (bs ! l) ! j" + using Scan\<^sub>F.prems items_nth_idem_bins_upd nth_idem_bins_upd length_map items_def by metis + finally show ?case + using Scan\<^sub>F.hyps by simp +next + case (Predict\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" + have "items (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1) ! l) ! j = items (?bs' ! l) ! j" + using Predict\<^sub>F.IH Predict\<^sub>F.prems length_nth_bin_bins_upd order.strict_trans2 items_def by (metis length_map) + also have "... = items (bs ! l) ! j" + using Predict\<^sub>F.prems items_nth_idem_bins_upd nth_idem_bins_upd length_map items_def by metis + finally show ?case + using Predict\<^sub>F.hyps by simp +qed simp_all + +lemma nth_bin_sub_Earley\<^sub>L_bin': + assumes "(k, \, \, bs) \ wf_earley_input" + shows "set (items (bs ! l)) \ set (items (Earley\<^sub>L_bin' k \ \ bs i ! l))" +proof standard + fix x + assume "x \ set (items (bs ! l))" + then obtain j where *: "j < length (items (bs ! l))" "items (bs ! l) ! j = x" + using in_set_conv_nth by metis + have "x = items (Earley\<^sub>L_bin' k \ \ bs i ! l) ! j" + using kth_Earley\<^sub>L_bin'_bins assms * by metis + moreover have "j < length (items (Earley\<^sub>L_bin' k \ \ bs i ! l))" + using assms *(1) length_nth_bin_Earley\<^sub>L_bin' less_le_trans by blast + ultimately show "x \ set (items (Earley\<^sub>L_bin' k \ \ bs i ! l))" + by simp +qed + +lemma nth_Earley\<^sub>L_bin'_eq: + assumes "(k, \, \, bs) \ wf_earley_input" + shows "l < k \ Earley\<^sub>L_bin' k \ \ bs i ! l = bs ! l" + by (induction i rule: Earley\<^sub>L_bin'_induct[OF assms]) (auto simp: bins_upd_def) + +lemma set_items_Earley\<^sub>L_bin'_eq: + assumes "(k, \, \, bs) \ wf_earley_input" + shows "l < k \ set (items (Earley\<^sub>L_bin' k \ \ bs i ! l)) = set (items (bs ! l))" + by (simp add: assms nth_Earley\<^sub>L_bin'_eq) + +lemma bins_upto_k0_Earley\<^sub>L_bin'_eq: + assumes "(k, \, \, bs) \ wf_earley_input" + shows "bins_upto (Earley\<^sub>L_bin k \ \ bs) k 0 = bins_upto bs k 0" + unfolding bins_upto_def bin_upto_def Earley\<^sub>L_bin_def using set_items_Earley\<^sub>L_bin'_eq assms nth_Earley\<^sub>L_bin'_eq by fastforce + +lemma wf_earley_input_Init\<^sub>L: + assumes "k \ length \" "wf_\ \" + shows "(k, \, \, Init\<^sub>L \ \) \ wf_earley_input" +proof - + let ?rs = "filter (\r. rule_head r = \ \) (\ \)" + let ?b0 = "map (\r. (Entry (init_item r 0) Null)) ?rs" + let ?bs = "replicate (length \ + 1) ([])" + have "distinct (items ?b0)" + using assms unfolding wf_bin_def wf_item_def wf_\_def distinct_rules_def items_def + by (auto simp: init_item_def distinct_map inj_on_def) + moreover have "\x \ set (items ?b0). wf_item \ \ x \ item_end x = 0" + using assms unfolding wf_bin_def wf_item_def by (auto simp: init_item_def items_def) + moreover have "wf_bins \ \ ?bs" + unfolding wf_bins_def wf_bin_def wf_bin_items_def items_def using less_Suc_eq_0_disj by force + ultimately show ?thesis + using assms length_replicate wf_earley_input_intro + unfolding wf_bin_def Init\<^sub>L_def wf_bin_def wf_bin_items_def wf_bins_def + by (metis (no_types, lifting) length_list_update nth_list_update_eq nth_list_update_neq) +qed + +lemma length_bins_Init\<^sub>L[simp]: + "length (Init\<^sub>L \ \) = length \ + 1" + by (simp add: Init\<^sub>L_def) + +lemma wf_earley_input_Earley\<^sub>L_bins[simp]: + assumes "k \ length \" "wf_\ \" + shows "(k, \, \, Earley\<^sub>L_bins k \ \) \ wf_earley_input" + using assms +proof (induction k) + case 0 + have "(k, \, \, Init\<^sub>L \ \) \ wf_earley_input" + using assms wf_earley_input_Init\<^sub>L by blast + thus ?case + by (simp add: assms(2) wf_earley_input_Init\<^sub>L wf_earley_input_Earley\<^sub>L_bin) +next + case (Suc k) + have "(Suc k, \, \, Earley\<^sub>L_bins k \ \) \ wf_earley_input" + using Suc.IH Suc.prems(1) Suc_leD assms(2) wf_earley_input_elim wf_earley_input_intro by metis + thus ?case + by (simp add: wf_earley_input_Earley\<^sub>L_bin) +qed + +lemma length_Earley\<^sub>L_bins[simp]: + assumes "k \ length \" "wf_\ \" + shows "length (Earley\<^sub>L_bins k \ \) = length (Init\<^sub>L \ \)" + using assms wf_earley_input_Earley\<^sub>L_bins wf_earley_input_elim by fastforce + +lemma wf_bins_Earley\<^sub>L_bins[simp]: + assumes "k \ length \" "wf_\ \" + shows "wf_bins \ \ (Earley\<^sub>L_bins k \ \)" + using assms wf_earley_input_Earley\<^sub>L_bins wf_earley_input_elim by fastforce + +lemma wf_bins_Earley\<^sub>L: + "wf_\ \ \ wf_bins \ \ (Earley\<^sub>L \ \)" + by (simp add: Earley\<^sub>L_def) + + +subsection \Soundness\ + +lemma Init\<^sub>L_eq_Init\<^sub>F: + "bins (Init\<^sub>L \ \) = Init\<^sub>F \" +proof - + let ?rs = "filter (\r. rule_head r = \ \) (\ \)" + let ?b0 = "map (\r. (Entry (init_item r 0) Null)) ?rs" + let ?bs = "replicate (length \ + 1) ([])" + have "bins (?bs[0 := ?b0]) = set (items ?b0)" + proof - + have "bins (?bs[0 := ?b0]) = \ {set (items ((?bs[0 := ?b0]) ! k)) |k. k < length (?bs[0 := ?b0])}" + unfolding bins_def by blast + also have "... = set (items ((?bs[0 := ?b0]) ! 0)) \ \ {set (items ((?bs[0 := ?b0]) ! k)) |k. k < length (?bs[0 := ?b0]) \ k \ 0}" + by fastforce + also have "... = set (items (?b0))" + by (auto simp: items_def) + finally show ?thesis . + qed + also have "... = Init\<^sub>F \" + by (auto simp: Init\<^sub>F_def items_def rule_head_def) + finally show ?thesis + by (auto simp: Init\<^sub>L_def) +qed + +lemma Scan\<^sub>L_sub_Scan\<^sub>F: + assumes "wf_bins \ \ bs" "bins bs \ I" "x \ set (items (bs ! k))" "k < length bs" "k < length \" + assumes "next_symbol x = Some a" + shows "set (items (Scan\<^sub>L k \ a x pre)) \ Scan\<^sub>F k \ I" +proof standard + fix y + assume *: "y \ set (items (Scan\<^sub>L k \ a x pre))" + have "x \ bin I k" + using kth_bin_sub_bins assms(1-4) items_def wf_bin_def wf_bins_def wf_bin_items_def bin_def by fastforce + { + assume #: "k < length \" "\!k = a" + hence "y = inc_item x (k+1)" + using * unfolding Scan\<^sub>L_def by (simp add: items_def) + hence "y \ Scan\<^sub>F k \ I" + using \x \ bin I k\ # assms(6) unfolding Scan\<^sub>F_def by blast + } + thus "y \ Scan\<^sub>F k \ I" + using * assms(5) unfolding Scan\<^sub>L_def by (auto simp: items_def) +qed + +lemma Predict\<^sub>L_sub_Predict\<^sub>F: + assumes "wf_bins \ \ bs" "bins bs \ I" "x \ set (items (bs ! k))" "k < length bs" + assumes "next_symbol x = Some X" + shows "set (items (Predict\<^sub>L k \ X)) \ Predict\<^sub>F k \ I" +proof standard + fix y + assume *: "y \ set (items (Predict\<^sub>L k \ X))" + have "x \ bin I k" + using kth_bin_sub_bins assms(1-4) items_def wf_bin_def wf_bins_def bin_def wf_bin_items_def by fast + let ?rs = "filter (\r. rule_head r = X) (\ \)" + let ?xs = "map (\r. init_item r k) ?rs" + have "y \ set ?xs" + using * unfolding Predict\<^sub>L_def items_def by simp + then obtain r where "y = init_item r k" "rule_head r = X" "r \ set (\ \)" "next_symbol x = Some (rule_head r)" + using assms(5) by auto + thus "y \ Predict\<^sub>F k \ I" + unfolding Predict\<^sub>F_def using \x \ bin I k\ by blast +qed + +lemma Complete\<^sub>L_sub_Complete\<^sub>F: + assumes "wf_bins \ \ bs" "bins bs \ I" "y \ set (items (bs ! k))" "k < length bs" + assumes "next_symbol y = None" + shows "set (items (Complete\<^sub>L k y bs red)) \ Complete\<^sub>F k I" +proof standard + fix x + assume *: "x \ set (items (Complete\<^sub>L k y bs red))" + have "y \ bin I k" + using kth_bin_sub_bins assms items_def wf_bin_def wf_bins_def bin_def wf_bin_items_def by fast + let ?orig = "bs ! item_origin y" + let ?xs = "filter_with_index (\x. next_symbol x = Some (item_rule_head y)) (items ?orig)" + let ?xs' = "map (\(x, pre). (Entry (inc_item x k) (PreRed (item_origin y, pre, red) []))) ?xs" + have 0: "item_origin y < length bs" + using wf_bins_def wf_bin_def wf_item_def wf_bin_items_def assms(1,3,4) + by (metis Orderings.preorder_class.dual_order.strict_trans1 leD not_le_imp_less) + { + fix z + assume *: "z \ set (map fst ?xs)" + have "next_symbol z = Some (item_rule_head y)" + using * by (simp add: filter_with_index_cong_filter) + moreover have "z \ bin I (item_origin y)" + using 0 * assms(1,2) bin_def kth_bin_sub_bins wf_bins_kth_bin filter_with_index_cong_filter + by (metis (mono_tags, lifting) filter_is_subset in_mono mem_Collect_eq) + ultimately have "next_symbol z = Some (item_rule_head y)" "z \ bin I (item_origin y)" + by simp_all + } + hence 1: "\z \ set (map fst ?xs). next_symbol z = Some (item_rule_head y) \ z \ bin I (item_origin y)" + by blast + obtain z where z: "x = inc_item z k" "z \ set (map fst ?xs)" + using * unfolding Complete\<^sub>L_def by (auto simp: rev_image_eqI items_def) + moreover have "next_symbol z = Some (item_rule_head y)" "z \ bin I (item_origin y)" + using 1 z by blast+ + ultimately show "x \ Complete\<^sub>F k I" + using \y \ bin I k\ assms(5) unfolding Complete\<^sub>F_def next_symbol_def by (auto split: if_splits) +qed + +lemma sound_Scan\<^sub>L: + assumes "wf_bins \ \ bs" "bins bs \ I" "x \ set (items (bs!k))" "k < length bs" "k < length \" + assumes "next_symbol x = Some a" "\x \ I. wf_item \ \ x" "\x \ I. sound_item \ \ x" + shows "\x \ set (items (Scan\<^sub>L k \ a x i)). sound_item \ \ x" +proof standard + fix y + assume "y \ set (items (Scan\<^sub>L k \ a x i))" + hence "y \ Scan\<^sub>F k \ I" + by (meson Scan\<^sub>L_sub_Scan\<^sub>F assms(1-6) in_mono) + thus "sound_item \ \ y" + using sound_Scan assms(7,8) unfolding Scan\<^sub>F_def inc_item_def bin_def + by (smt (verit, best) item.exhaust_sel mem_Collect_eq) +qed + +lemma sound_Predict\<^sub>L: + assumes "wf_bins \ \ bs" "bins bs \ I" "x \ set (items (bs!k))" "k < length bs" + assumes "next_symbol x = Some X" "\x \ I. wf_item \ \ x" "\x \ I. sound_item \ \ x" + shows "\x \ set (items (Predict\<^sub>L k \ X)). sound_item \ \ x" +proof standard + fix y + assume "y \ set (items (Predict\<^sub>L k \ X))" + hence "y \ Predict\<^sub>F k \ I" + by (meson Predict\<^sub>L_sub_Predict\<^sub>F assms(1-5) subsetD) + thus "sound_item \ \ y" + using sound_Predict assms(6,7) unfolding Predict\<^sub>F_def init_item_def bin_def + by (smt (verit, del_insts) item.exhaust_sel mem_Collect_eq) +qed + +lemma sound_Complete\<^sub>L: + assumes "wf_bins \ \ bs" "bins bs \ I" "y \ set (items (bs!k))" "k < length bs" + assumes "next_symbol y = None" "\x \ I. wf_item \ \ x" "\x \ I. sound_item \ \ x" + shows "\x \ set (items (Complete\<^sub>L k y bs i)). sound_item \ \ x" +proof standard + fix x + assume "x \ set (items (Complete\<^sub>L k y bs i))" + hence "x \ Complete\<^sub>F k I" + using Complete\<^sub>L_sub_Complete\<^sub>F assms(1-5) by blast + thus "sound_item \ \ x" + using sound_Complete assms(6,7) unfolding Complete\<^sub>F_def inc_item_def bin_def + by (smt (verit, del_insts) item.exhaust_sel mem_Collect_eq) +qed + +lemma sound_Earley\<^sub>L_bin': + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "\x \ bins bs. sound_item \ \ x" + shows "\x \ bins (Earley\<^sub>L_bin' k \ \ bs i). sound_item \ \ x" + using assms +proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) + case (Complete\<^sub>F k \ \ bs i x) + let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" + have "x \ set (items (bs ! k))" + using Complete\<^sub>F.hyps(1,2) by force + hence "\x \ set (items (Complete\<^sub>L k x bs i)). sound_item \ \ x" + using sound_Complete\<^sub>L Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items by fastforce + moreover have "(k, \, \, ?bs') \ wf_earley_input" + using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast + ultimately have "\x \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1)). sound_item \ \ x" + using Complete\<^sub>F.IH Complete\<^sub>F.prems(2) length_bins_upd bins_bins_upd wf_earley_input_elim + Suc_eq_plus1 Un_iff le_imp_less_Suc by metis + thus ?case + using Complete\<^sub>F.hyps by simp +next + case (Scan\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" + have "x \ set (items (bs ! k))" + using Scan\<^sub>F.hyps(1,2) by force + hence "\x \ set (items (Scan\<^sub>L k \ a x i)). sound_item \ \ x" + using sound_Scan\<^sub>L Scan\<^sub>F.hyps(3,5) Scan\<^sub>F.prems(1,2) wf_earley_input_elim wf_bins_impl_wf_items by fast + moreover have "(k, \, \, ?bs') \ wf_earley_input" + using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis + ultimately have "\x \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1)). sound_item \ \ x" + using Scan\<^sub>F.IH Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(2) length_bins_upd bins_bins_upd wf_earley_input_elim + by (metis UnE add_less_cancel_right) + thus ?case + using Scan\<^sub>F.hyps by simp +next + case (Predict\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" + have "x \ set (items (bs ! k))" + using Predict\<^sub>F.hyps(1,2) by force + hence "\x \ set (items(Predict\<^sub>L k \ a)). sound_item \ \ x" + using sound_Predict\<^sub>L Predict\<^sub>F.hyps(3) Predict\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items by fast + moreover have "(k, \, \, ?bs') \ wf_earley_input" + using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis + ultimately have "\x \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1)). sound_item \ \ x" + using Predict\<^sub>F.IH Predict\<^sub>F.prems(2) length_bins_upd bins_bins_upd wf_earley_input_elim + by (metis Suc_eq_plus1 UnE) + thus ?case + using Predict\<^sub>F.hyps by simp +qed simp_all + +lemma sound_Earley\<^sub>L_bin: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "\x \ bins bs. sound_item \ \ x" + shows "\x \ bins (Earley\<^sub>L_bin k \ \ bs). sound_item \ \ x" + using sound_Earley\<^sub>L_bin' assms Earley\<^sub>L_bin_def by metis + +lemma Earley\<^sub>L_bin'_sub_Earley\<^sub>F_bin: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "bins bs \ I" + shows "bins (Earley\<^sub>L_bin' k \ \ bs i) \ Earley\<^sub>F_bin k \ \ I" + using assms +proof (induction i arbitrary: I rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) + case (Base k \ \ bs i) + thus ?case + using Earley\<^sub>F_bin_mono by fastforce +next + case (Complete\<^sub>F k \ \ bs i x) + let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" + have "x \ set (items (bs ! k))" + using Complete\<^sub>F.hyps(1,2) by force + hence "bins ?bs' \ I \ Complete\<^sub>F k I" + using Complete\<^sub>L_sub_Complete\<^sub>F Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems(1,2) bins_bins_upd wf_earley_input_elim by blast + moreover have "(k, \, \, ?bs') \ wf_earley_input" + using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast + ultimately have "bins (Earley\<^sub>L_bin' k \ \ bs i) \ Earley\<^sub>F_bin k \ \ (I \ Complete\<^sub>F k I)" + using Complete\<^sub>F.IH Complete\<^sub>F.hyps by simp + also have "... \ Earley\<^sub>F_bin k \ \ (Earley\<^sub>F_bin k \ \ I)" + using Complete\<^sub>F_Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_sub_mono by (metis Un_subset_iff) + finally show ?case + using Earley\<^sub>F_bin_idem by blast +next + case (Scan\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" + have "x \ set (items (bs ! k))" + using Scan\<^sub>F.hyps(1,2) by force + hence "bins ?bs' \ I \ Scan\<^sub>F k \ I" + using Scan\<^sub>L_sub_Scan\<^sub>F Scan\<^sub>F.hyps(3,5) Scan\<^sub>F.prems bins_bins_upd wf_earley_input_elim + by (metis add_mono1 sup_mono) + moreover have "(k, \, \, ?bs') \ wf_earley_input" + using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis + ultimately have "bins (Earley\<^sub>L_bin' k \ \ bs i) \ Earley\<^sub>F_bin k \ \ (I \ Scan\<^sub>F k \ I)" + using Scan\<^sub>F.IH Scan\<^sub>F.hyps by simp + thus ?case + using Scan\<^sub>F_Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_sub_mono Earley\<^sub>F_bin_idem by (metis le_supI order_trans) +next + case (Pass k \ \ bs i x a) + thus ?case + by simp +next + case (Predict\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" + have "x \ set (items (bs ! k))" + using Predict\<^sub>F.hyps(1,2) by force + hence "bins ?bs' \ I \ Predict\<^sub>F k \ I" + using Predict\<^sub>L_sub_Predict\<^sub>F Predict\<^sub>F.hyps(3) Predict\<^sub>F.prems bins_bins_upd wf_earley_input_elim + by (metis sup_mono) + moreover have "(k, \, \, ?bs') \ wf_earley_input" + using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis + ultimately have "bins (Earley\<^sub>L_bin' k \ \ bs i) \ Earley\<^sub>F_bin k \ \ (I \ Predict\<^sub>F k \ I)" + using Predict\<^sub>F.IH Predict\<^sub>F.hyps by simp + thus ?case + using Predict\<^sub>F_Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_mono Earley\<^sub>F_bin_sub_mono Earley\<^sub>F_bin_idem by (metis le_supI order_trans) +qed + +lemma Earley\<^sub>L_bin_sub_Earley\<^sub>F_bin: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "bins bs \ I" + shows "bins (Earley\<^sub>L_bin k \ \ bs) \ Earley\<^sub>F_bin k \ \ I" + using assms Earley\<^sub>L_bin'_sub_Earley\<^sub>F_bin Earley\<^sub>L_bin_def by metis + +lemma Earley\<^sub>L_bins_sub_Earley\<^sub>F_bins: + assumes "k \ length \" "wf_\ \" + shows "bins (Earley\<^sub>L_bins k \ \) \ Earley\<^sub>F_bins k \ \" + using assms +proof (induction k) + case 0 + have "(k, \, \, Init\<^sub>L \ \) \ wf_earley_input" + using assms(1) assms(2) wf_earley_input_Init\<^sub>L by blast + thus ?case + by (simp add: Init\<^sub>L_eq_Init\<^sub>F Earley\<^sub>L_bin_sub_Earley\<^sub>F_bin assms(2) wf_earley_input_Init\<^sub>L) +next + case (Suc k) + have "(Suc k, \, \, Earley\<^sub>L_bins k \ \) \ wf_earley_input" + by (simp add: Suc.prems(1) Suc_leD assms(2) wf_earley_input_intro) + thus ?case + by (simp add: Suc.IH Suc.prems(1) Suc_leD Earley\<^sub>L_bin_sub_Earley\<^sub>F_bin assms(2)) +qed + +lemma Earley\<^sub>L_sub_Earley\<^sub>F: + "wf_\ \ \ bins (Earley\<^sub>L \ \) \ Earley\<^sub>F \ \" + using Earley\<^sub>L_bins_sub_Earley\<^sub>F_bins Earley\<^sub>F_def Earley\<^sub>L_def by (metis dual_order.refl) + +theorem soundness_Earley\<^sub>L: + assumes "wf_\ \" "recognizing (bins (Earley\<^sub>L \ \)) \ \" + shows "derives \ [\ \] \" + using assms Earley\<^sub>L_sub_Earley\<^sub>F recognizing_def soundness_Earley\<^sub>F by (meson subsetD) + + +subsection \Completeness\ + +lemma bin_bins_upto_bins_eq: + assumes "wf_bins \ \ bs" "k < length bs" "i \ length (items (bs ! k))" "l \ k" + shows "bin (bins_upto bs k i) l = bin (bins bs) l" + unfolding bins_upto_def bins_def bin_def using assms nat_less_le + apply (auto simp: nth_list_update bin_upto_eq_set_items wf_bins_kth_bin items_def) + apply (metis imageI nle_le order_trans, fast) + done + +lemma impossible_complete_item: + assumes "wf_\ \" "wf_item \ \ x" "sound_item \ \ x" + assumes "is_complete x" "item_origin x = k" "item_end x = k" "nonempty_derives \" + shows False +proof - + have "derives \ [item_rule_head x] []" + using assms(3-6) by (simp add: slice_empty is_complete_def sound_item_def item_\_def ) + moreover have "is_nonterminal \ (item_rule_head x)" + using assms(1,2) unfolding wf_item_def item_rule_head_def rule_head_def + by (metis prod.collapse rule_nonterminal_type) + ultimately show ?thesis + using assms(7) nonempty_derives_def is_nonterminal_def by metis +qed + +lemma Complete\<^sub>F_Un_eq_terminal: + assumes "next_symbol z = Some a" "is_terminal \ a" "\x \ I. wf_item \ \ x" "wf_item \ \ z" "wf_\ \" + shows "Complete\<^sub>F k (I \ {z}) = Complete\<^sub>F k I" +proof (rule ccontr) + assume "Complete\<^sub>F k (I \ {z}) \ Complete\<^sub>F k I" + hence "Complete\<^sub>F k I \ Complete\<^sub>F k (I \ {z})" + using Complete\<^sub>F_sub_mono by blast + then obtain w x y where *: + "w \ Complete\<^sub>F k (I \ {z})" "w \ Complete\<^sub>F k I" "w = inc_item x k" + "x \ bin (I \ {z}) (item_origin y)" "y \ bin (I \ {z}) k" + "is_complete y" "next_symbol x = Some (item_rule_head y)" + unfolding Complete\<^sub>F_def by fast + show False + proof (cases "x = z") + case True + have "is_nonterminal \ (item_rule_head y)" + using *(5,6) assms(1,3-5) + apply (clarsimp simp: wf_item_def bin_def item_rule_head_def rule_head_def next_symbol_def) + by (metis prod.exhaust_sel rule_nonterminal_type) + thus ?thesis + using True *(7) assms(1,2,5) is_terminal_nonterminal by fastforce + next + case False + thus ?thesis + using * assms(1) by (auto simp: next_symbol_def Complete\<^sub>F_def bin_def) + qed +qed + +lemma Complete\<^sub>F_Un_eq_nonterminal: + assumes "wf_\ \" "\x \ I. wf_item \ \ x" "\x \ I. sound_item \ \ x" + assumes "nonempty_derives \" "wf_item \ \ z" + assumes "item_end z = k" "next_symbol z \ None" + shows "Complete\<^sub>F k (I \ {z}) = Complete\<^sub>F k I" +proof (rule ccontr) + assume "Complete\<^sub>F k (I \ {z}) \ Complete\<^sub>F k I" + hence "Complete\<^sub>F k I \ Complete\<^sub>F k (I \ {z})" + using Complete\<^sub>F_sub_mono by blast + then obtain x x' y where *: + "x \ Complete\<^sub>F k (I \ {z})" "x \ Complete\<^sub>F k I" "x = inc_item x' k" + "x' \ bin (I \ {z}) (item_origin y)" "y \ bin (I \ {z}) k" + "is_complete y" "next_symbol x' = Some (item_rule_head y)" + unfolding Complete\<^sub>F_def by fast + consider (A) "x' = z" | (B) "y = z" + using *(2-7) Complete\<^sub>F_def by (auto simp: bin_def; blast) + thus False + proof cases + case A + have "item_origin y = k" + using *(4) A bin_def assms(6) by (metis (mono_tags, lifting) mem_Collect_eq) + moreover have "item_end y = k" + using *(5) bin_def by blast + moreover have "sound_item \ \ y" + using *(5,6) assms(3,7) by (auto simp: bin_def next_symbol_def sound_item_def) + moreover have "wf_item \ \ y" + using *(5) assms(2,5) wf_item_def by (auto simp: bin_def) + ultimately show ?thesis + using impossible_complete_item *(6) assms(1,4) by blast + next + case B + thus ?thesis + using *(6) assms(7) by (auto simp: next_symbol_def) + qed +qed + +lemma wf_item_in_kth_bin: + "wf_bins \ \ bs \ x \ bins bs \ item_end x = k \ x \ set (items (bs ! k))" + using bins_bin_exists wf_bins_kth_bin wf_bins_def by blast + +lemma Complete\<^sub>F_bins_upto_eq_bins: + assumes "wf_bins \ \ bs" "k < length bs" "i \ length (items (bs ! k))" + shows "Complete\<^sub>F k (bins_upto bs k i) = Complete\<^sub>F k (bins bs)" +proof - + have "\l. l \ k \ bin (bins_upto bs k i) l = bin (bins bs) l" + using bin_bins_upto_bins_eq[OF assms] by blast + moreover have "\x \ bins bs. wf_item \ \ x" + using assms(1) wf_bins_impl_wf_items by metis + ultimately show ?thesis + unfolding Complete\<^sub>F_def bin_def wf_item_def wf_item_def by auto +qed + +lemma Complete\<^sub>F_sub_bins_Un_Complete\<^sub>L: + assumes "Complete\<^sub>F k I \ bins bs" "I \ bins bs" "is_complete z" "wf_bins \ \ bs" "wf_item \ \ z" + shows "Complete\<^sub>F k (I \ {z}) \ bins bs \ set (items (Complete\<^sub>L k z bs red))" +proof standard + fix w + assume "w \ Complete\<^sub>F k (I \ {z})" + then obtain x y where *: + "w = inc_item x k" "x \ bin (I \ {z}) (item_origin y)" "y \ bin (I \ {z}) k" + "is_complete y" "next_symbol x = Some (item_rule_head y)" + unfolding Complete\<^sub>F_def by blast + consider (A) "x = z" | (B) "y = z" | "\ (x = z \ y = z)" + by blast + thus "w \ bins bs \ set (items (Complete\<^sub>L k z bs red))" + proof cases + case A + thus ?thesis + using *(5) assms(3) by (auto simp: next_symbol_def) + next + case B + let ?orig = "bs ! item_origin z" + let ?is = "filter_with_index (\x. next_symbol x = Some (item_rule_head z)) (items ?orig)" + have "x \ bin I (item_origin y)" + using B *(2) *(5) assms(3) by (auto simp: next_symbol_def bin_def) + moreover have "bin I (item_origin z) \ set (items (bs ! item_origin z))" + using wf_item_in_kth_bin assms(2,4) bin_def by blast + ultimately have "x \ set (map fst ?is)" + using *(5) B by (simp add: filter_with_index_cong_filter in_mono) + thus ?thesis + unfolding Complete\<^sub>L_def *(1) by (auto simp: rev_image_eqI items_def) + next + case 3 + thus ?thesis + using * assms(1) Complete\<^sub>F_def by (auto simp: bin_def; blast) + qed +qed + +lemma Complete\<^sub>L_eq_item_origin: + "bs ! item_origin y = bs' ! item_origin y \ Complete\<^sub>L k y bs red = Complete\<^sub>L k y bs' red" + by (auto simp: Complete\<^sub>L_def) + +lemma kth_bin_bins_upto_empty: + assumes "wf_bins \ \ bs" "k < length bs" + shows "bin (bins_upto bs k 0) k = {}" +proof - + { + fix x + assume "x \ bins_upto bs k 0" + then obtain l where "x \ set (items (bs ! l))" "l < k" + unfolding bins_upto_def bin_upto_def by blast + hence "item_end x = l" + using wf_bins_kth_bin assms by fastforce + hence "item_end x < k" + using \l < k\ by blast + } + thus ?thesis + by (auto simp: bin_def) +qed + +lemma Earley\<^sub>L_bin'_mono: + assumes "(k, \, \, bs) \ wf_earley_input" + shows "bins bs \ bins (Earley\<^sub>L_bin' k \ \ bs i)" + using assms +proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) + case (Complete\<^sub>F k \ \ bs i x) + let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" + have wf: "(k, \, \, ?bs') \ wf_earley_input" + using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast + hence "bins bs \ bins ?bs'" + using length_bins_upd bins_bins_upd wf_earley_input_elim by (metis Un_upper1) + also have "... \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" + using wf Complete\<^sub>F.IH by blast + finally show ?case + using Complete\<^sub>F.hyps by simp +next + case (Scan\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" + have wf: "(k, \, \, ?bs') \ wf_earley_input" + using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis + hence "bins bs \ bins ?bs'" + using Scan\<^sub>F.hyps(5) length_bins_upd bins_bins_upd wf_earley_input_elim + by (metis add_mono1 sup_ge1) + also have "... \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" + using wf Scan\<^sub>F.IH by blast + finally show ?case + using Scan\<^sub>F.hyps by simp +next + case (Predict\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" + have wf: "(k, \, \, ?bs') \ wf_earley_input" + using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis + hence "bins bs \ bins ?bs'" + using length_bins_upd bins_bins_upd wf_earley_input_elim by (metis sup_ge1) + also have "... \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" + using wf Predict\<^sub>F.IH by blast + finally show ?case + using Predict\<^sub>F.hyps by simp +qed simp_all + +lemma Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin': + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "Earley\<^sub>F_bin_step k \ \ (bins_upto bs k i) \ bins bs" + assumes "\x \ bins bs. sound_item \ \ x" "is_word \ \" "nonempty_derives \" + shows "Earley\<^sub>F_bin_step k \ \ (bins bs) \ bins (Earley\<^sub>L_bin' k \ \ bs i)" + using assms +proof (induction i rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) + case (Base k \ \ bs i) + have "bin (bins bs) k = bin (bins_upto bs k i) k" + using Base.hyps Base.prems(1) bin_bins_upto_bins_eq wf_earley_input_elim by blast + thus ?case + using Scan\<^sub>F_bin_absorb Predict\<^sub>F_bin_absorb Complete\<^sub>F_bins_upto_eq_bins wf_earley_input_elim + Base.hyps Base.prems(1,2,3,5) Earley\<^sub>F_bin_step_def Complete\<^sub>F_Earley\<^sub>F_bin_step_mono + Predict\<^sub>F_Earley\<^sub>F_bin_step_mono Scan\<^sub>F_Earley\<^sub>F_bin_step_mono Earley\<^sub>L_bin'_mono + by (metis (no_types, lifting) Un_assoc sup.orderE) +next + case (Complete\<^sub>F k \ \ bs i x) + let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" + have x: "x \ set (items (bs ! k))" + using Complete\<^sub>F.hyps(1,2) by auto + have wf: "(k, \, \, ?bs') \ wf_earley_input" + using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast + hence sound: "\x \ set (items (Complete\<^sub>L k x bs i)). sound_item \ \ x" + using sound_Complete\<^sub>L Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items x + by (metis dual_order.refl) + have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" + proof - + have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Scan\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" + using Complete\<^sub>F.hyps(1) bins_upto_Suc_Un length_nth_bin_bins_upd items_def + by (metis length_map linorder_not_less sup.boundedE sup.order_iff) + also have "... = Scan\<^sub>F k \ (bins_upto bs k i \ {x})" + using Complete\<^sub>F.hyps(1,2) Complete\<^sub>F.prems(1) items_nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim + by (metis dual_order.refl items_def length_map not_le_imp_less) + also have "... \ bins bs \ Scan\<^sub>F k \ {x}" + using Complete\<^sub>F.prems(2,3) Scan\<^sub>F_Un Scan\<^sub>F_Earley\<^sub>F_bin_step_mono by fastforce + also have "... = bins bs" + using Complete\<^sub>F.hyps(3) by (auto simp: Scan\<^sub>F_def bin_def) + finally show ?thesis + using Complete\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by blast + qed + moreover have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" + proof - + have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Predict\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" + using Complete\<^sub>F.hyps(1) bins_upto_Suc_Un length_nth_bin_bins_upd + by (metis dual_order.strict_trans1 items_def length_map not_le_imp_less) + also have "... = Predict\<^sub>F k \ (bins_upto bs k i \ {x})" + using Complete\<^sub>F.hyps(1,2) Complete\<^sub>F.prems(1) items_nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim + by (metis dual_order.refl items_def length_map not_le_imp_less) + also have "... \ bins bs \ Predict\<^sub>F k \ {x}" + using Complete\<^sub>F.prems(2,3) Predict\<^sub>F_Un Predict\<^sub>F_Earley\<^sub>F_bin_step_mono by blast + also have "... = bins bs" + using Complete\<^sub>F.hyps(3) by (auto simp: Predict\<^sub>F_def bin_def) + finally show ?thesis + using Complete\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by blast + qed + moreover have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) \ bins ?bs'" + proof - + have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) = Complete\<^sub>F k (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" + using bins_upto_Suc_Un length_nth_bin_bins_upd Complete\<^sub>F.hyps(1) + by (metis (no_types, opaque_lifting) dual_order.trans items_def length_map not_le_imp_less) + also have "... = Complete\<^sub>F k (bins_upto bs k i \ {x})" + using items_nth_idem_bins_upd Complete\<^sub>F.hyps(1,2) bins_upto_kth_nth_idem Complete\<^sub>F.prems(1) wf_earley_input_elim + by (metis dual_order.refl items_def length_map not_le_imp_less) + also have "... \ bins bs \ set (items (Complete\<^sub>L k x bs i))" + using Complete\<^sub>F_sub_bins_Un_Complete\<^sub>L Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems(1,2,3) next_symbol_def + bins_upto_sub_bins wf_bins_kth_bin x Complete\<^sub>F_Earley\<^sub>F_bin_step_mono wf_earley_input_elim + by (smt (verit, best) option.distinct(1) subset_trans) + finally show ?thesis + using Complete\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by blast + qed + ultimately have "Earley\<^sub>F_bin_step k \ \ (bins ?bs') \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" + using Complete\<^sub>F.IH Complete\<^sub>F.prems sound wf Earley\<^sub>F_bin_step_def bins_upto_sub_bins + wf_earley_input_elim bins_bins_upd + by (metis UnE sup.boundedI) + thus ?case + using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) Earley\<^sub>L_bin'_simps(2) Earley\<^sub>F_bin_step_sub_mono bins_bins_upd wf_earley_input_elim + by (smt (verit, best) sup.coboundedI2 sup.orderE sup_ge1) +next + case (Scan\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" + have x: "x \ set (items (bs ! k))" + using Scan\<^sub>F.hyps(1,2) by auto + hence sound: "\x \ set (items (Scan\<^sub>L k \ a x i)). sound_item \ \ x" + using sound_Scan\<^sub>L Scan\<^sub>F.hyps(3,5) Scan\<^sub>F.prems(1,2,3) wf_earley_input_elim wf_bins_impl_wf_items x + by (metis dual_order.refl) + have wf: "(k, \, \, ?bs') \ wf_earley_input" + using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis + have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" + proof - + have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Scan\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" + using bins_upto_Suc_Un Scan\<^sub>F.hyps(1) nth_idem_bins_upd + by (metis Suc_eq_plus1 items_def length_map lessI less_not_refl not_le_imp_less) + also have "... = Scan\<^sub>F k \ (bins_upto bs k i \ {x})" + using Scan\<^sub>F.hyps(1,2,5) Scan\<^sub>F.prems(1,2) nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim + by (metis add_mono_thms_linordered_field(1) items_def length_map less_add_one linorder_le_less_linear not_add_less1) + also have "... \ bins bs \ Scan\<^sub>F k \ {x}" + using Scan\<^sub>F.prems(2,3) Scan\<^sub>F_Un Scan\<^sub>F_Earley\<^sub>F_bin_step_mono by fastforce + finally have *: "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins bs \ Scan\<^sub>F k \ {x}" . + show ?thesis + proof cases + assume a1: "\!k = a" + hence "Scan\<^sub>F k \ {x} = {inc_item x (k+1)}" + using Scan\<^sub>F.hyps(1-3,5) Scan\<^sub>F.prems(1,2) wf_earley_input_elim apply (auto simp: Scan\<^sub>F_def bin_def) + using wf_bins_kth_bin x by blast + hence "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins bs \ {inc_item x (k+1)}" + using * by blast + also have "... = bins bs \ set (items (Scan\<^sub>L k \ a x i))" + using a1 Scan\<^sub>F.hyps(5) by (auto simp: Scan\<^sub>L_def items_def) + also have "... = bins ?bs'" + using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by (metis add_mono1) + finally show ?thesis . + next + assume a1: "\ \!k = a" + hence "Scan\<^sub>F k \ {x} = {}" + using Scan\<^sub>F.hyps(3) by (auto simp: Scan\<^sub>F_def bin_def) + hence "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins bs" + using * by blast + also have "... \ bins ?bs'" + using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd + by (metis Un_left_absorb add_strict_right_mono subset_Un_eq) + finally show ?thesis . + qed + qed + moreover have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" + proof - + have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Predict\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" + using bins_upto_Suc_Un Scan\<^sub>F.hyps(1) nth_idem_bins_upd + by (metis Suc_eq_plus1 dual_order.refl items_def length_map lessI linorder_not_less) + also have "... = Predict\<^sub>F k \ (bins_upto bs k i \ {x})" + using Scan\<^sub>F.hyps(1,2,5) Scan\<^sub>F.prems(1,2) nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim + by (metis add_strict_right_mono items_def le_add1 length_map less_add_one linorder_not_le) + also have "... \ bins bs \ Predict\<^sub>F k \ {x}" + using Scan\<^sub>F.prems(2,3) Predict\<^sub>F_Un Predict\<^sub>F_Earley\<^sub>F_bin_step_mono by fastforce + also have "... = bins bs" + using Scan\<^sub>F.hyps(3,4) Scan\<^sub>F.prems(1) is_terminal_nonterminal wf_earley_input_elim + by (auto simp: Predict\<^sub>F_def bin_def rule_head_def, fastforce) + finally show ?thesis + using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1) by (simp add: bins_bins_upd sup.coboundedI1 wf_earley_input_elim) + qed + moreover have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) \ bins ?bs'" + proof - + have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) = Complete\<^sub>F k (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" + using bins_upto_Suc_Un Scan\<^sub>F.hyps(1) nth_idem_bins_upd + by (metis Suc_eq_plus1 items_def length_map lessI less_not_refl not_le_imp_less) + also have "... = Complete\<^sub>F k (bins_upto bs k i \ {x})" + using Scan\<^sub>F.hyps(1,2,5) Scan\<^sub>F.prems(1,2) nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim + by (metis add_mono1 items_def length_map less_add_one linorder_not_le not_add_less1) + also have "... = Complete\<^sub>F k (bins_upto bs k i)" + using Complete\<^sub>F_Un_eq_terminal Scan\<^sub>F.hyps(3,4) Scan\<^sub>F.prems bins_upto_sub_bins subset_iff + wf_bins_impl_wf_items wf_bins_kth_bin wf_item_def x wf_earley_input_elim + by (smt (verit, ccfv_threshold)) + finally show ?thesis + using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1,2,3) Complete\<^sub>F_Earley\<^sub>F_bin_step_mono by (auto simp: bins_bins_upd wf_earley_input_elim, blast) + qed + ultimately have "Earley\<^sub>F_bin_step k \ \ (bins ?bs') \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" + using Scan\<^sub>F.IH Scan\<^sub>F.prems Scan\<^sub>F.hyps(5) sound wf Earley\<^sub>F_bin_step_def bins_upto_sub_bins wf_earley_input_elim + bins_bins_upd by (metis UnE add_mono1 le_supI) + thus ?case + using Earley\<^sub>F_bin_step_sub_mono Earley\<^sub>L_bin'_simps(3) Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd + by (smt (verit, ccfv_SIG) add_mono1 sup.cobounded1 sup.coboundedI2 sup.orderE) +next + case (Pass k \ \ bs i x a) + have x: "x \ set (items (bs ! k))" + using Pass.hyps(1,2) by auto + have "Scan\<^sub>F k \ (bins_upto bs k (i + 1)) \ bins bs" + using Scan\<^sub>F_def Pass.hyps(5) by auto + moreover have "Predict\<^sub>F k \ (bins_upto bs k (i + 1)) \ bins bs" + proof - + have "Predict\<^sub>F k \ (bins_upto bs k (i + 1)) = Predict\<^sub>F k \ (bins_upto bs k i \ {items (bs ! k) ! i})" + using bins_upto_Suc_Un Pass.hyps(1) by (metis items_def length_map not_le_imp_less) + also have "... = Predict\<^sub>F k \ (bins_upto bs k i \ {x})" + using Pass.hyps(1,2,5) nth_idem_bins_upd bins_upto_kth_nth_idem by simp + also have "... \ bins bs \ Predict\<^sub>F k \ {x}" + using Pass.prems(2) Predict\<^sub>F_Un Predict\<^sub>F_Earley\<^sub>F_bin_step_mono by blast + also have "... = bins bs" + using Pass.hyps(3,4) Pass.prems(1) is_terminal_nonterminal wf_earley_input_elim + by (auto simp: Predict\<^sub>F_def bin_def rule_head_def, fastforce) + finally show ?thesis + using bins_bins_upd Pass.hyps(5) Pass.prems(3) by auto + qed + moreover have "Complete\<^sub>F k (bins_upto bs k (i + 1)) \ bins bs" + proof - + have "Complete\<^sub>F k (bins_upto bs k (i + 1)) = Complete\<^sub>F k (bins_upto bs k i \ {x})" + using bins_upto_Suc_Un Pass.hyps(1,2) + by (metis items_def length_map not_le_imp_less) + also have "... = Complete\<^sub>F k (bins_upto bs k i)" + using Complete\<^sub>F_Un_eq_terminal Pass.hyps Pass.prems bins_upto_sub_bins subset_iff + wf_bins_impl_wf_items wf_item_def wf_bins_kth_bin x wf_earley_input_elim by (smt (verit, best)) + finally show ?thesis + using Pass.prems(1,2) Complete\<^sub>F_Earley\<^sub>F_bin_step_mono wf_earley_input_elim by blast + qed + ultimately have "Earley\<^sub>F_bin_step k \ \ (bins bs) \ bins (Earley\<^sub>L_bin' k \ \ bs (i+1))" + using Pass.IH Pass.prems Earley\<^sub>F_bin_step_def bins_upto_sub_bins wf_earley_input_elim + by (metis le_sup_iff) + thus ?case + using bins_bins_upd Pass.hyps Pass.prems by simp +next + case (Predict\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" + have "k \ length \ \ \ \!k = a" + using Predict\<^sub>F.hyps(4) Predict\<^sub>F.prems(4) is_word_is_terminal leI by blast + have x: "x \ set (items (bs ! k))" + using Predict\<^sub>F.hyps(1,2) by auto + hence sound: "\x \ set (items(Predict\<^sub>L k \ a)). sound_item \ \ x" + using sound_Predict\<^sub>L Predict\<^sub>F.hyps(3) Predict\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items by fast + have wf: "(k, \, \, ?bs') \ wf_earley_input" + using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis + have len: "i < length (items (?bs' ! k))" + using length_nth_bin_bins_upd Predict\<^sub>F.hyps(1) + by (metis dual_order.strict_trans1 items_def length_map linorder_not_less) + have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" + proof - + have "Scan\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Scan\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" + using Predict\<^sub>F.hyps(1) bins_upto_Suc_Un by (metis items_def len length_map) + also have "... = Scan\<^sub>F k \ (bins_upto bs k i \ {x})" + using Predict\<^sub>F.hyps(1,2) Predict\<^sub>F.prems(1) items_nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim + by (metis dual_order.refl items_def length_map not_le_imp_less) + also have "... \ bins bs \ Scan\<^sub>F k \ {x}" + using Predict\<^sub>F.prems(2,3) Scan\<^sub>F_Un Scan\<^sub>F_Earley\<^sub>F_bin_step_mono by fastforce + also have "... = bins bs" + using Predict\<^sub>F.hyps(3) \length \ \ k \ \ ! k \ a\ by (auto simp: Scan\<^sub>F_def bin_def) + finally show ?thesis + using Predict\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by blast + qed + moreover have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) \ bins ?bs'" + proof - + have "Predict\<^sub>F k \ (bins_upto ?bs' k (i + 1)) = Predict\<^sub>F k \ (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" + using Predict\<^sub>F.hyps(1) bins_upto_Suc_Un by (metis items_def len length_map) + also have "... = Predict\<^sub>F k \ (bins_upto bs k i \ {x})" + using Predict\<^sub>F.hyps(1,2) Predict\<^sub>F.prems(1) items_nth_idem_bins_upd bins_upto_kth_nth_idem wf_earley_input_elim + by (metis dual_order.refl items_def length_map not_le_imp_less) + also have "... \ bins bs \ Predict\<^sub>F k \ {x}" + using Predict\<^sub>F.prems(2,3) Predict\<^sub>F_Un Predict\<^sub>F_Earley\<^sub>F_bin_step_mono by fastforce + also have "... = bins bs \ set (items (Predict\<^sub>L k \ a))" + using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1-3) wf_earley_input_elim + apply (auto simp: Predict\<^sub>F_def Predict\<^sub>L_def bin_def items_def) + using wf_bins_kth_bin x by blast + finally show ?thesis + using Predict\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by blast + qed + moreover have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) \ bins ?bs'" + proof - + have "Complete\<^sub>F k (bins_upto ?bs' k (i + 1)) = Complete\<^sub>F k (bins_upto ?bs' k i \ {items (?bs' ! k) ! i})" + using bins_upto_Suc_Un len by (metis items_def length_map) + also have "... = Complete\<^sub>F k (bins_upto bs k i \ {x})" + using items_nth_idem_bins_upd Predict\<^sub>F.hyps(1,2) Predict\<^sub>F.prems(1) bins_upto_kth_nth_idem wf_earley_input_elim + by (metis dual_order.refl items_def length_map not_le_imp_less) + also have "... = Complete\<^sub>F k (bins_upto bs k i)" + using Complete\<^sub>F_Un_eq_nonterminal Predict\<^sub>F.prems bins_upto_sub_bins Predict\<^sub>F.hyps(3) + subset_eq wf_bins_kth_bin x wf_bins_impl_wf_items wf_item_def wf_earley_input_elim + by (smt (verit, ccfv_SIG) option.simps(3)) + also have "... \ bins bs" + using Complete\<^sub>F_Earley\<^sub>F_bin_step_mono Predict\<^sub>F.prems(2) by blast + finally show ?thesis + using bins_bins_upd Predict\<^sub>F.prems(1,2,3) wf_earley_input_elim by (metis Un_upper1 dual_order.trans) + qed + ultimately have "Earley\<^sub>F_bin_step k \ \ (bins ?bs') \ bins (Earley\<^sub>L_bin' k \ \ ?bs' (i+1))" + using Predict\<^sub>F.IH Predict\<^sub>F.prems sound wf Earley\<^sub>F_bin_step_def bins_upto_sub_bins + bins_bins_upd wf_earley_input_elim by (metis UnE le_supI) + hence "Earley\<^sub>F_bin_step k \ \ (bins ?bs') \ bins (Earley\<^sub>L_bin' k \ \ bs i)" + using Predict\<^sub>F.hyps Earley\<^sub>L_bin'_simps(5) by simp + moreover have "Earley\<^sub>F_bin_step k \ \ (bins bs) \ Earley\<^sub>F_bin_step k \ \ (bins ?bs')" + using Earley\<^sub>F_bin_step_sub_mono Predict\<^sub>F.prems(1) wf_earley_input_elim bins_bins_upd by (metis Un_upper1) + ultimately show ?case + by blast +qed + +lemma Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "Earley\<^sub>F_bin_step k \ \ (bins_upto bs k 0) \ bins bs" + assumes "\x \ bins bs. sound_item \ \ x" "is_word \ \" "nonempty_derives \" + shows "Earley\<^sub>F_bin_step k \ \ (bins bs) \ bins (Earley\<^sub>L_bin k \ \ bs)" + using assms Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin' Earley\<^sub>L_bin_def by metis + +lemma bins_eq_items_Complete\<^sub>L: + assumes "bins_eq_items as bs" "item_origin x < length as" + shows "items (Complete\<^sub>L k x as i) = items (Complete\<^sub>L k x bs i)" +proof - + let ?orig_a = "as ! item_origin x" + let ?orig_b = "bs ! item_origin x" + have "items ?orig_a = items ?orig_b" + using assms by (metis (no_types, opaque_lifting) bins_eq_items_def length_map nth_map) + thus ?thesis + unfolding Complete\<^sub>L_def by simp +qed + +lemma Earley\<^sub>L_bin'_bins_eq: + assumes "(k, \, \, as) \ wf_earley_input" + assumes "bins_eq_items as bs" "wf_bins \ \ as" + shows "bins_eq_items (Earley\<^sub>L_bin' k \ \ as i) (Earley\<^sub>L_bin' k \ \ bs i)" + using assms +proof (induction i arbitrary: bs rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) + case (Base k \ \ as i) + have "Earley\<^sub>L_bin' k \ \ as i = as" + by (simp add: Base.hyps) + moreover have "Earley\<^sub>L_bin' k \ \ bs i = bs" + using Base.hyps Base.prems(1,2) unfolding bins_eq_items_def + by (metis Earley\<^sub>L_bin'_simps(1) length_map nth_map wf_earley_input_elim) + ultimately show ?case + using Base.prems(2) by presburger +next + case (Complete\<^sub>F k \ \ as i x) + let ?as' = "bins_upd as k (Complete\<^sub>L k x as i)" + let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" + have k: "k < length as" + using Complete\<^sub>F.prems(1) wf_earley_input_elim by blast + hence wf_x: "wf_item \ \ x" + using Complete\<^sub>F.hyps(1,2) Complete\<^sub>F.prems(3) wf_bins_kth_bin by fastforce + have "(k, \, \, ?as') \ wf_earley_input" + using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast + moreover have "bins_eq_items ?as' ?bs'" + using Complete\<^sub>F.hyps(1,2) Complete\<^sub>F.prems(2,3) bins_eq_items_dist_bins_upd bins_eq_items_Complete\<^sub>L + k wf_x wf_bins_kth_bin wf_item_def by (metis dual_order.strict_trans2 leI nth_mem) + ultimately have "bins_eq_items (Earley\<^sub>L_bin' k \ \ ?as' (i + 1)) (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" + using Complete\<^sub>F.IH wf_earley_input_elim by blast + moreover have "Earley\<^sub>L_bin' k \ \ as i = Earley\<^sub>L_bin' k \ \ ?as' (i+1)" + using Complete\<^sub>F.hyps by simp + moreover have "Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" + using Complete\<^sub>F.hyps Complete\<^sub>F.prems unfolding bins_eq_items_def + by (metis Earley\<^sub>L_bin'_simps(2) map_eq_imp_length_eq nth_map wf_earley_input_elim) + ultimately show ?case + by argo +next + case (Scan\<^sub>F k \ \ as i x a) + let ?as' = "bins_upd as (k+1) (Scan\<^sub>L k \ a x i)" + let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" + have "(k, \, \, ?as') \ wf_earley_input" + using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by fast + moreover have "bins_eq_items ?as' ?bs'" + using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1,2) bins_eq_items_dist_bins_upd add_mono1 wf_earley_input_elim by metis + ultimately have "bins_eq_items (Earley\<^sub>L_bin' k \ \ ?as' (i + 1)) (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" + using Scan\<^sub>F.IH wf_earley_input_elim by blast + moreover have "Earley\<^sub>L_bin' k \ \ as i = Earley\<^sub>L_bin' k \ \ ?as' (i+1)" + using Scan\<^sub>F.hyps by simp + moreover have "Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" + using Scan\<^sub>F.hyps Scan\<^sub>F.prems unfolding bins_eq_items_def + by (smt (verit, ccfv_threshold) Earley\<^sub>L_bin'_simps(3) length_map nth_map wf_earley_input_elim) + ultimately show ?case + by argo +next + case (Pass k \ \ as i x a) + have "bins_eq_items (Earley\<^sub>L_bin' k \ \ as (i + 1)) (Earley\<^sub>L_bin' k \ \ bs (i + 1))" + using Pass.prems Pass.IH by blast + moreover have "Earley\<^sub>L_bin' k \ \ as i = Earley\<^sub>L_bin' k \ \ as (i+1)" + using Pass.hyps by simp + moreover have "Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ bs (i+1)" + using Pass.hyps Pass.prems unfolding bins_eq_items_def + by (metis Earley\<^sub>L_bin'_simps(4) map_eq_imp_length_eq nth_map wf_earley_input_elim) + ultimately show ?case + by argo +next + case (Predict\<^sub>F k \ \ as i x a) + let ?as' = "bins_upd as k (Predict\<^sub>L k \ a)" + let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" + have "(k, \, \, ?as') \ wf_earley_input" + using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by fast + moreover have "bins_eq_items ?as' ?bs'" + using Predict\<^sub>F.prems(1,2) bins_eq_items_dist_bins_upd wf_earley_input_elim by blast + ultimately have "bins_eq_items (Earley\<^sub>L_bin' k \ \ ?as' (i + 1)) (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" + using Predict\<^sub>F.IH wf_earley_input_elim by blast + moreover have "Earley\<^sub>L_bin' k \ \ as i = Earley\<^sub>L_bin' k \ \ ?as' (i+1)" + using Predict\<^sub>F.hyps by simp + moreover have "Earley\<^sub>L_bin' k \ \ bs i = Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" + using Predict\<^sub>F.hyps Predict\<^sub>F.prems unfolding bins_eq_items_def + by (metis Earley\<^sub>L_bin'_simps(5) length_map nth_map wf_earley_input_elim) + ultimately show ?case + by argo +qed + +lemma Earley\<^sub>L_bin'_idem: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "i \ j" "\x \ bins bs. sound_item \ \ x" "nonempty_derives \" + shows "bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ bs i) j) = bins (Earley\<^sub>L_bin' k \ \ bs i)" + using assms +proof (induction i arbitrary: j rule: Earley\<^sub>L_bin'_induct[OF assms(1), case_names Base Complete\<^sub>F Scan\<^sub>F Pass Predict\<^sub>F]) + case (Complete\<^sub>F k \ \ bs i x) + let ?bs' = "bins_upd bs k (Complete\<^sub>L k x bs i)" + have x: "x \ set (items (bs ! k))" + using Complete\<^sub>F.hyps(1,2) by auto + have wf: "(k, \, \, ?bs') \ wf_earley_input" + using Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) wf_earley_input_Complete\<^sub>L by blast + hence "\x \ set (items (Complete\<^sub>L k x bs i)). sound_item \ \ x" + using sound_Complete\<^sub>L Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items x + by (metis dual_order.refl) + hence sound: "\x \ bins ?bs'. sound_item \ \ x" + by (metis Complete\<^sub>F.prems(1,3) UnE bins_bins_upd wf_earley_input_elim) + show ?case + proof cases + assume "i+1 \ j" + thus ?thesis + using wf sound Complete\<^sub>F Earley\<^sub>L_bin'_simps(2) by metis + next + assume "\ i+1 \ j" + hence "i = j" + using Complete\<^sub>F.prems(2) by simp + have "bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ bs i) j) = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) j)" + using Earley\<^sub>L_bin'_simps(2) Complete\<^sub>F.hyps(1-3) by simp + also have "... = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) (j+1))" + proof - + let ?bs'' = "Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" + have "length (items (?bs'' ! k)) \ length (items (bs ! k))" + using length_nth_bin_Earley\<^sub>L_bin' length_nth_bin_bins_upd order_trans wf Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) + by (smt (verit, ccfv_threshold) Earley\<^sub>L_bin'_simps(2)) + hence 0: "\ length (items (?bs'' ! k)) \ j" + using \i = j\ Complete\<^sub>F.hyps(1) by linarith + have "x = items (?bs' ! k) ! j" + using \i = j\ items_nth_idem_bins_upd Complete\<^sub>F.hyps(1,2) + by (metis items_def length_map not_le_imp_less) + hence 1: "x = items (?bs'' ! k) ! j" + using \i = j\ kth_Earley\<^sub>L_bin'_bins Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) Earley\<^sub>L_bin'_simps(2) leI by metis + have "bins (Earley\<^sub>L_bin' k \ \ ?bs'' j) = bins (Earley\<^sub>L_bin' k \ \ (bins_upd ?bs'' k (Complete\<^sub>L k x ?bs'' i)) (j+1))" + using Earley\<^sub>L_bin'_simps(2) 0 1 Complete\<^sub>F.hyps(1,3) Complete\<^sub>F.prems(2) \i = j\ by auto + moreover have "bins_eq_items (bins_upd ?bs'' k (Complete\<^sub>L k x ?bs'' i)) ?bs''" + proof - + have "k < length bs" + using Complete\<^sub>F.prems(1) wf_earley_input_elim by blast + have 0: "set (Complete\<^sub>L k x bs i) = set (Complete\<^sub>L k x ?bs'' i)" + proof (cases "item_origin x = k") + case True + thus ?thesis + using impossible_complete_item kth_bin_sub_bins Complete\<^sub>F.hyps(3) Complete\<^sub>F.prems wf_earley_input_elim + wf_bins_kth_bin x next_symbol_def by (metis option.distinct(1) subsetD) + next + case False + hence "item_origin x < k" + using x Complete\<^sub>F.prems(1) wf_bins_kth_bin wf_item_def nat_less_le by (metis wf_earley_input_elim) + hence "bs ! item_origin x = ?bs'' ! item_origin x" + using False nth_idem_bins_upd nth_Earley\<^sub>L_bin'_eq wf by metis + thus ?thesis + using Complete\<^sub>L_eq_item_origin by metis + qed + have "set (items (Complete\<^sub>L k x bs i)) \ set (items (?bs' ! k))" + by (simp add: \k < length bs\ bins_upd_def set_items_bin_upds) + hence "set (items (Complete\<^sub>L k x ?bs'' i)) \ set (items (?bs' ! k))" + using 0 by (simp add: items_def) + also have "... \ set (items (?bs'' ! k))" + by (simp add: wf nth_bin_sub_Earley\<^sub>L_bin') + finally show ?thesis + using bins_eq_items_bins_upd by blast + qed + moreover have "(k, \, \, bins_upd ?bs'' k (Complete\<^sub>L k x ?bs'' i)) \ wf_earley_input" + using wf_earley_input_Earley\<^sub>L_bin' wf_earley_input_Complete\<^sub>L Complete\<^sub>F.hyps Complete\<^sub>F.prems(1) + \length (items (bs ! k)) \ length (items (?bs'' ! k))\ kth_Earley\<^sub>L_bin'_bins 0 1 by blast + ultimately show ?thesis + using Earley\<^sub>L_bin'_bins_eq bins_eq_items_imp_eq_bins wf_earley_input_elim by blast + qed + also have "... = bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" + using Complete\<^sub>F.IH[OF wf _ sound Complete\<^sub>F.prems(4)] \i = j\ by blast + finally show ?thesis + using Complete\<^sub>F.hyps by simp + qed +next + case (Scan\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs (k+1) (Scan\<^sub>L k \ a x i)" + have x: "x \ set (items (bs ! k))" + using Scan\<^sub>F.hyps(1,2) by auto + hence "\x \ set (items (Scan\<^sub>L k \ a x i)). sound_item \ \ x" + using sound_Scan\<^sub>L Scan\<^sub>F.hyps(3,5) Scan\<^sub>F.prems(1,2,3) wf_earley_input_elim wf_bins_impl_wf_items x + by (metis dual_order.refl) + hence sound: "\x \ bins ?bs'. sound_item \ \ x" + using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems(1,3) bins_bins_upd wf_earley_input_elim + by (metis UnE add_less_cancel_right) + have wf: "(k, \, \, ?bs') \ wf_earley_input" + using Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) wf_earley_input_Scan\<^sub>L by metis + show ?case + proof cases + assume "i+1 \ j" + thus ?thesis + using sound Scan\<^sub>F by (metis Earley\<^sub>L_bin'_simps(3) wf_earley_input_Scan\<^sub>L) + next + assume "\ i+1 \ j" + hence "i = j" + using Scan\<^sub>F.prems(2) by auto + have "bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ bs i) j) = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) j)" + using Scan\<^sub>F.hyps by simp + also have "... = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) (j+1))" + proof - + let ?bs'' = "Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" + have "length (items (?bs'' ! k)) \ length (items (bs ! k))" + using length_nth_bin_Earley\<^sub>L_bin' length_nth_bin_bins_upd order_trans Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) Earley\<^sub>L_bin'_simps(3) + by (smt (verit, ccfv_SIG)) + hence "bins (Earley\<^sub>L_bin' k \ \ ?bs'' j) = bins (Earley\<^sub>L_bin' k \ \ (bins_upd ?bs'' (k+1) (Scan\<^sub>L k \ a x i)) (j+1))" + using \i = j\ kth_Earley\<^sub>L_bin'_bins nth_idem_bins_upd Earley\<^sub>L_bin'_simps(3) Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) by (smt (verit, best) leI le_trans) + moreover have "bins_eq_items (bins_upd ?bs'' (k+1) (Scan\<^sub>L k \ a x i)) ?bs''" + proof - + have "k+1 < length bs" + using Scan\<^sub>F.hyps(5) Scan\<^sub>F.prems wf_earley_input_elim by fastforce+ + hence "set (items (Scan\<^sub>L k \ a x i)) \ set (items (?bs' ! (k+1)))" + by (simp add: bins_upd_def set_items_bin_upds) + also have "... \ set (items (?bs'' ! (k+1)))" + using wf nth_bin_sub_Earley\<^sub>L_bin' by blast + finally show ?thesis + using bins_eq_items_bins_upd by blast + qed + moreover have "(k, \, \, bins_upd ?bs'' (k+1) (Scan\<^sub>L k \ a x i)) \ wf_earley_input" + using wf_earley_input_Earley\<^sub>L_bin' wf_earley_input_Scan\<^sub>L Scan\<^sub>F.hyps Scan\<^sub>F.prems(1) + \length (items (bs ! k)) \ length (items (?bs'' ! k))\ kth_Earley\<^sub>L_bin'_bins + by (smt (verit, ccfv_SIG) Earley\<^sub>L_bin'_simps(3) linorder_not_le order.trans) + ultimately show ?thesis + using Earley\<^sub>L_bin'_bins_eq bins_eq_items_imp_eq_bins wf_earley_input_elim by blast + qed + also have "... = bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" + using \i = j\ Scan\<^sub>F.IH Scan\<^sub>F.prems Scan\<^sub>F.hyps sound wf_earley_input_Scan\<^sub>L by fast + finally show ?thesis + using Scan\<^sub>F.hyps by simp + qed +next + case (Pass k \ \ bs i x a) + show ?case + proof cases + assume "i+1 \ j" + thus ?thesis + using Pass by (metis Earley\<^sub>L_bin'_simps(4)) + next + assume "\ i+1 \ j" + show ?thesis + using Pass Earley\<^sub>L_bin'_simps(1,4) kth_Earley\<^sub>L_bin'_bins by (metis Suc_eq_plus1 Suc_leI antisym_conv2 not_le_imp_less) + qed +next + case (Predict\<^sub>F k \ \ bs i x a) + let ?bs' = "bins_upd bs k (Predict\<^sub>L k \ a)" + have x: "x \ set (items (bs ! k))" + using Predict\<^sub>F.hyps(1,2) by auto + hence "\x \ set (items(Predict\<^sub>L k \ a)). sound_item \ \ x" + using sound_Predict\<^sub>L Predict\<^sub>F.hyps(3) Predict\<^sub>F.prems wf_earley_input_elim wf_bins_impl_wf_items by fast + hence sound: "\x \ bins ?bs'. sound_item \ \ x" + using Predict\<^sub>F.prems(1,3) UnE bins_bins_upd wf_earley_input_elim by metis + have wf: "(k, \, \, ?bs') \ wf_earley_input" + using Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) wf_earley_input_Predict\<^sub>L by metis + have len: "i < length (items (?bs' ! k))" + using length_nth_bin_bins_upd Predict\<^sub>F.hyps(1) Orderings.preorder_class.dual_order.strict_trans1 linorder_not_less + by (metis items_def length_map) + show ?case + proof cases + assume "i+1 \ j" + thus ?thesis + using sound wf Predict\<^sub>F by (metis Earley\<^sub>L_bin'_simps(5)) + next + assume "\ i+1 \ j" + hence "i = j" + using Predict\<^sub>F.prems(2) by auto + have "bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ bs i) j) = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) j)" + using Predict\<^sub>F.hyps by simp + also have "... = bins (Earley\<^sub>L_bin' k \ \ (Earley\<^sub>L_bin' k \ \ ?bs' (i+1)) (j+1))" + proof - + let ?bs'' = "Earley\<^sub>L_bin' k \ \ ?bs' (i+1)" + have "length (items (?bs'' ! k)) \ length (items (bs ! k))" + using length_nth_bin_Earley\<^sub>L_bin' length_nth_bin_bins_upd order_trans wf + by (metis (no_types, lifting) items_def length_map) + hence "bins (Earley\<^sub>L_bin' k \ \ ?bs'' j) = bins (Earley\<^sub>L_bin' k \ \ (bins_upd ?bs'' k (Predict\<^sub>L k \ a)) (j+1))" + using \i = j\ kth_Earley\<^sub>L_bin'_bins nth_idem_bins_upd Earley\<^sub>L_bin'_simps(5) Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) length_bins_Earley\<^sub>L_bin' + wf_bins_Earley\<^sub>L_bin' wf_bins_kth_bin wf_item_def x by (smt (verit, ccfv_SIG) linorder_not_le order.trans) + moreover have "bins_eq_items (bins_upd ?bs'' k (Predict\<^sub>L k \ a)) ?bs''" + proof - + have "k < length bs" + using wf_earley_input_elim[OF Predict\<^sub>F.prems(1)] by blast + hence "set (items (Predict\<^sub>L k \ a)) \ set (items (?bs' ! k))" + by (simp add: bins_upd_def set_items_bin_upds) + also have "... \ set (items (?bs'' ! k))" + using wf nth_bin_sub_Earley\<^sub>L_bin' by blast + finally show ?thesis + using bins_eq_items_bins_upd by blast + qed + moreover have "(k, \, \, bins_upd ?bs'' k (Predict\<^sub>L k \ a)) \ wf_earley_input" + using wf_earley_input_Earley\<^sub>L_bin' wf_earley_input_Predict\<^sub>L Predict\<^sub>F.hyps Predict\<^sub>F.prems(1) + \length (items (bs ! k)) \ length (items (?bs'' ! k))\ kth_Earley\<^sub>L_bin'_bins + by (smt (verit, best) Earley\<^sub>L_bin'_simps(5) dual_order.trans not_le_imp_less) + ultimately show ?thesis + using Earley\<^sub>L_bin'_bins_eq bins_eq_items_imp_eq_bins wf_earley_input_elim by blast + qed + also have "... = bins (Earley\<^sub>L_bin' k \ \ ?bs' (i + 1))" + using \i = j\ Predict\<^sub>F.IH Predict\<^sub>F.prems sound wf by (metis order_refl) + finally show ?thesis + using Predict\<^sub>F.hyps by simp + qed +qed simp + +lemma Earley\<^sub>L_bin_idem: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "\x \ bins bs. sound_item \ \ x" "nonempty_derives \" + shows "bins (Earley\<^sub>L_bin k \ \ (Earley\<^sub>L_bin k \ \ bs)) = bins (Earley\<^sub>L_bin k \ \ bs)" + using assms Earley\<^sub>L_bin'_idem Earley\<^sub>L_bin_def le0 by metis + +lemma funpower_Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "Earley\<^sub>F_bin_step k \ \ (bins_upto bs k 0) \ bins bs" "\x \ bins bs. sound_item \ \ x" + assumes "is_word \ \" "nonempty_derives \" + shows "funpower (Earley\<^sub>F_bin_step k \ \) n (bins bs) \ bins (Earley\<^sub>L_bin k \ \ bs)" + using assms +proof (induction n) + case 0 + thus ?case + using Earley\<^sub>L_bin'_mono Earley\<^sub>L_bin_def by (simp add: Earley\<^sub>L_bin'_mono Earley\<^sub>L_bin_def) +next + case (Suc n) + have 0: "Earley\<^sub>F_bin_step k \ \ (bins_upto (Earley\<^sub>L_bin k \ \ bs) k 0) \ bins (Earley\<^sub>L_bin k \ \ bs)" + using Earley\<^sub>L_bin'_mono bins_upto_k0_Earley\<^sub>L_bin'_eq assms(1,2) Earley\<^sub>L_bin_def order_trans + by (metis (no_types, lifting)) + have "funpower (Earley\<^sub>F_bin_step k \ \) (Suc n) (bins bs) \ Earley\<^sub>F_bin_step k \ \ (bins (Earley\<^sub>L_bin k \ \ bs))" + using Earley\<^sub>F_bin_step_sub_mono Suc by (metis funpower.simps(2)) + also have "... \ bins (Earley\<^sub>L_bin k \ \ (Earley\<^sub>L_bin k \ \ bs))" + using Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin Suc.prems wf_bins_Earley\<^sub>L_bin sound_Earley\<^sub>L_bin 0 wf_earley_input_Earley\<^sub>L_bin by blast + also have "... \ bins (Earley\<^sub>L_bin k \ \ bs)" + using Earley\<^sub>L_bin_idem Suc.prems by blast + finally show ?case . +qed + +lemma Earley\<^sub>F_bin_sub_Earley\<^sub>L_bin: + assumes "(k, \, \, bs) \ wf_earley_input" + assumes "Earley\<^sub>F_bin_step k \ \ (bins_upto bs k 0) \ bins bs" "\x \ bins bs. sound_item \ \ x" + assumes "is_word \ \" "nonempty_derives \" + shows "Earley\<^sub>F_bin k \ \ (bins bs) \ bins (Earley\<^sub>L_bin k \ \ bs)" + using assms funpower_Earley\<^sub>F_bin_step_sub_Earley\<^sub>L_bin Earley\<^sub>F_bin_def elem_limit_simp by fastforce + +lemma Earley\<^sub>F_bins_sub_Earley\<^sub>L_bins: + assumes "k \ length \" "wf_\ \" + assumes "is_word \ \" "nonempty_derives \" + shows "Earley\<^sub>F_bins k \ \ \ bins (Earley\<^sub>L_bins k \ \)" + using assms +proof (induction k) + case 0 + hence "Earley\<^sub>F_bin 0 \ \ (Init\<^sub>F \) \ bins (Earley\<^sub>L_bin 0 \ \ (Init\<^sub>L \ \))" + using Earley\<^sub>F_bin_sub_Earley\<^sub>L_bin Init\<^sub>L_eq_Init\<^sub>F length_bins_Init\<^sub>L Init\<^sub>L_eq_Init\<^sub>F sound_Init bins_upto_empty + Earley\<^sub>F_bin_step_empty bins_upto_sub_bins wf_earley_input_Init\<^sub>L wf_earley_input_elim + by (smt (verit, ccfv_threshold) Init\<^sub>F_sub_Earley basic_trans_rules(31) sound_Earley wf_bins_impl_wf_items) + thus ?case + by simp +next + case (Suc k) + have wf: "(Suc k, \, \, Earley\<^sub>L_bins k \ \) \ wf_earley_input" + by (simp add: Suc.prems(1) Suc_leD assms(2) wf_earley_input_intro) + have sub: "Earley\<^sub>F_bin_step (Suc k) \ \ (bins_upto (Earley\<^sub>L_bins k \ \) (Suc k) 0) \ bins (Earley\<^sub>L_bins k \ \)" + proof - + have "bin (bins_upto (Earley\<^sub>L_bins k \ \) (Suc k) 0) (Suc k) = {}" + using kth_bin_bins_upto_empty wf Suc.prems wf_earley_input_elim by blast + hence "Earley\<^sub>F_bin_step (Suc k) \ \ (bins_upto (Earley\<^sub>L_bins k \ \) (Suc k) 0) = bins_upto (Earley\<^sub>L_bins k \ \) (Suc k) 0" + unfolding Earley\<^sub>F_bin_step_def Scan\<^sub>F_def Complete\<^sub>F_def Predict\<^sub>F_def bin_def by blast + also have "... \ bins (Earley\<^sub>L_bins k \ \)" + using wf Suc.prems bins_upto_sub_bins wf_earley_input_elim by blast + finally show ?thesis . + qed + have sound: "\x \ bins (Earley\<^sub>L_bins k \ \). sound_item \ \ x" + using Suc Earley\<^sub>L_bins_sub_Earley\<^sub>F_bins by (metis Suc_leD Earley\<^sub>F_bins_sub_Earley in_mono sound_Earley wf_Earley) + have "Earley\<^sub>F_bins (Suc k) \ \ \ Earley\<^sub>F_bin (Suc k) \ \ (bins (Earley\<^sub>L_bins k \ \))" + using Suc Earley\<^sub>F_bin_sub_mono by simp + also have "... \ bins (Earley\<^sub>L_bin (Suc k) \ \ (Earley\<^sub>L_bins k \ \))" + using Earley\<^sub>F_bin_sub_Earley\<^sub>L_bin wf sub sound Suc.prems by fastforce + finally show ?case + by simp +qed + +lemma Earley\<^sub>F_sub_Earley\<^sub>L: + assumes "wf_\ \" "is_word \ \" "nonempty_derives \" + shows "Earley\<^sub>F \ \ \ bins (Earley\<^sub>L \ \)" + using assms Earley\<^sub>F_bins_sub_Earley\<^sub>L_bins Earley\<^sub>F_def Earley\<^sub>L_def by (metis le_refl) + +theorem completeness_Earley\<^sub>L: + assumes "derives \ [\ \] \" "is_word \ \" "wf_\ \" "nonempty_derives \" + shows "recognizing (bins (Earley\<^sub>L \ \)) \ \" + using assms Earley\<^sub>F_sub_Earley\<^sub>L Earley\<^sub>L_sub_Earley\<^sub>F completeness_Earley\<^sub>F by (metis subset_antisym) + + +subsection \Correctness\ + +theorem Earley_eq_Earley\<^sub>L: + assumes "wf_\ \" "is_word \ \" "nonempty_derives \" + shows "Earley \ \ = bins (Earley\<^sub>L \ \)" + using assms Earley\<^sub>F_sub_Earley\<^sub>L Earley\<^sub>L_sub_Earley\<^sub>F Earley_eq_Earley\<^sub>F by blast + +theorem correctness_Earley\<^sub>L: + assumes "wf_\ \" "is_word \ \" "nonempty_derives \" + shows "recognizing (bins (Earley\<^sub>L \ \)) \ \ \ derives \ [\ \] \" + using assms Earley_eq_Earley\<^sub>L correctness_Earley by fastforce + +end diff --git a/thys/Earley_Parser/Examples.thy b/thys/Earley_Parser/Examples.thy new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/Examples.thy @@ -0,0 +1,196 @@ +theory Examples + imports Earley_Parser +begin + +section \Epsilon productions\ + +definition \_free :: "'a cfg \ bool" where + "\_free \ \ (\r \ set (\ \). rule_body r \ [])" + +lemma \_free_impl_non_empty_sentence_deriv: + "\_free \ \ a \ [] \ \ Derivation \ a D []" +proof (induction "length D" arbitrary: a D rule: nat_less_induct) + case 1 + show ?case + proof (rule ccontr) + assume assm: "\ \ Derivation \ a D []" + show False + proof (cases "D = []") + case True + then show ?thesis + using "1.prems"(2) assm by auto + next + case False + then obtain d D' \ where *: + "D = d # D'" "Derives1 \ a (fst d) (snd d) \" "Derivation \ \ D' []" "snd d \ set (\ \)" + using list.exhaust assm Derives1_def by (metis Derivation.simps(2)) + show ?thesis + proof cases + assume "\ = []" + thus ?thesis + using *(2,4) Derives1_split \_free_def rule_body_def "1.prems"(1) by (metis append_is_Nil_conv) + next + assume "\ \ = []" + thus ?thesis + using *(1,3) "1.hyps" "1.prems"(1) by auto + qed + qed + qed +qed + +lemma \_free_impl_non_empty_deriv: + "\_free \ \ \N \ set (\ \). \ derives \ [N] []" + using \_free_impl_non_empty_sentence_deriv derives_implies_Derivation by (metis not_Cons_self2) + +lemma nonempty_deriv_impl_\_free: + assumes "\N \ set (\ \). \ derives \ [N] []" "\(N, \) \ set (\ \). N \ set (\ \)" + shows "\_free \" +proof (rule ccontr) + assume "\ \_free \" + then obtain N \ where *: "(N, \) \ set (\ \)" "rule_body (N, \) = []" + unfolding \_free_def by auto + hence "derives1 \ [N] []" + unfolding derives1_def rule_body_def by auto + hence "derives \ [N] []" + by auto + moreover have "N \ set (\ \)" + using *(1) assms(2) by blast + ultimately show False + using assms(1) by blast +qed + +lemma nonempty_deriv_iff_\_free: + assumes "\(N, \) \ set (\ \). N \ set (\ \)" + shows "(\N \ set (\ \). \ derives \ [N] []) \ \_free \" + using \_free_impl_non_empty_deriv nonempty_deriv_impl_\_free[OF _ assms] by blast + +section \Example 1: Addition\ + +datatype t1 = x | plus +datatype n1 = S +datatype s1 = Terminal t1 | Nonterminal n1 + +definition nonterminals1 :: "s1 list" where + "nonterminals1 = [Nonterminal S]" + +definition terminals1 :: "s1 list" where + "terminals1 = [Terminal x, Terminal plus]" + +definition rules1 :: "s1 rule list" where + "rules1 = [ + (Nonterminal S, [Terminal x]), + (Nonterminal S, [Nonterminal S, Terminal plus, Nonterminal S]) + ]" + +definition start_symbol1 :: s1 where + "start_symbol1 = Nonterminal S" + +definition cfg1 :: "s1 cfg" where + "cfg1 = CFG nonterminals1 terminals1 rules1 start_symbol1" + +definition inp1 :: "s1 list" where + "inp1 = [Terminal x, Terminal plus, Terminal x, Terminal plus, Terminal x]" + +lemmas cfg1_defs = cfg1_def nonterminals1_def terminals1_def rules1_def start_symbol1_def + +lemma wf_\1: + "wf_\ cfg1" + by (auto simp: wf_\_defs cfg1_defs) + +lemma is_word_inp1: + "is_word cfg1 inp1" + by (auto simp: is_word_def is_terminal_def cfg1_defs inp1_def) + +lemma nonempty_derives1: + "nonempty_derives cfg1" + by (auto simp: \_free_def cfg1_defs rule_body_def nonempty_derives_def \_free_impl_non_empty_deriv) + +lemma correctness1: + "recognizing (bins (Earley\<^sub>L cfg1 inp1)) cfg1 inp1 \ derives cfg1 [\ cfg1] inp1" + using correctness_Earley\<^sub>L wf_\1 is_word_inp1 nonempty_derives1 by blast + +lemma wf_tree1: + assumes "build_tree cfg1 inp1 (Earley\<^sub>L cfg1 inp1) = Some t" + shows "wf_rule_tree cfg1 t \ root_tree t = \ cfg1 \ yield_tree t = inp1" + using assms nonempty_derives1 wf_\1 wf_rule_root_yield_tree_build_tree_Earley\<^sub>L by blast + +lemma correctness_tree1: + "(\t. build_tree cfg1 inp1 (Earley\<^sub>L cfg1 inp1) = Some t) \ derives cfg1 [\ cfg1] inp1" + using correctness_build_tree_Earley\<^sub>L is_word_inp1 nonempty_derives1 wf_\1 by blast + +lemma wf_trees1: + assumes "build_trees cfg1 inp1 (Earley\<^sub>L cfg1 inp1) = Some fs" "f \ set fs" "t \ set (trees f)" + shows "wf_rule_tree cfg1 t \ root_tree t = \ cfg1 \ yield_tree t = inp1" + using assms nonempty_derives1 wf_\1 wf_rule_root_yield_tree_build_trees_Earley\<^sub>L by blast + +lemma soundness_trees1: + assumes "build_trees cfg1 inp1 (Earley\<^sub>L cfg1 inp1) = Some fs" "f \ set fs" "t \ set (trees f)" + shows "derives cfg1 [\ cfg1] inp1" + using assms is_word_inp1 nonempty_derives1 soundness_build_trees_Earley\<^sub>L wf_\1 by blast + +section \Example 2: Cyclic reduction pointers\ + +datatype t2 = x +datatype n2 = A | B +datatype s2 = Terminal t2 | Nonterminal n2 + +definition nonterminals2 :: "s2 list" where + "nonterminals2 = [Nonterminal A, Nonterminal B]" + +definition terminals2 :: "s2 list" where + "terminals2 = [Terminal x]" + +definition rules2 :: "s2 rule list" where + "rules2 = [ + (Nonterminal B, [Nonterminal A]), + (Nonterminal A, [Nonterminal B]), + (Nonterminal A, [Terminal x]) + ]" + +definition start_symbol2 :: s2 where + "start_symbol2 = Nonterminal A" + +definition cfg2 :: "s2 cfg" where + "cfg2 = CFG nonterminals2 terminals2 rules2 start_symbol2" + +definition inp2 :: "s2 list" where + "inp2 = [Terminal x]" + +lemmas cfg2_defs = cfg2_def nonterminals2_def terminals2_def rules2_def start_symbol2_def + +lemma wf_\2: + "wf_\ cfg2" + by (auto simp: wf_\_defs cfg2_defs) + +lemma is_word_inp2: + "is_word cfg2 inp2" + by (auto simp: is_word_def is_terminal_def cfg2_defs inp2_def) + +lemma nonempty_derives2: + "nonempty_derives cfg2" + by (auto simp: \_free_def cfg2_defs rule_body_def nonempty_derives_def \_free_impl_non_empty_deriv) + +lemma correctness2: + "recognizing (bins (Earley\<^sub>L cfg2 inp2)) cfg2 inp2 \ derives cfg2 [\ cfg2] inp2" + using correctness_Earley\<^sub>L wf_\2 is_word_inp2 nonempty_derives2 by blast + +lemma wf_tree2: + assumes "build_tree cfg2 inp2 (Earley\<^sub>L cfg2 inp2) = Some t" + shows "wf_rule_tree cfg2 t \ root_tree t = \ cfg2 \ yield_tree t = inp2" + using assms nonempty_derives2 wf_\2 wf_rule_root_yield_tree_build_tree_Earley\<^sub>L by blast + +lemma correctness_tree2: + "(\t. build_tree cfg2 inp2 (Earley\<^sub>L cfg2 inp2) = Some t) \ derives cfg2 [\ cfg2] inp2" + using correctness_build_tree_Earley\<^sub>L is_word_inp2 nonempty_derives2 wf_\2 by blast + +lemma wf_trees2: + assumes "build_trees cfg2 inp2 (Earley\<^sub>L cfg2 inp2) = Some fs" "f \ set fs" "t \ set (trees f)" + shows "wf_rule_tree cfg2 t \ root_tree t = \ cfg2 \ yield_tree t = inp2" + using assms nonempty_derives2 wf_\2 wf_rule_root_yield_tree_build_trees_Earley\<^sub>L by blast + +lemma soundness_trees2: + assumes "build_trees cfg2 inp2 (Earley\<^sub>L cfg2 inp2) = Some fs" "f \ set fs" "t \ set (trees f)" + shows "derives cfg2 [\ cfg2] inp2" + using assms is_word_inp2 nonempty_derives2 soundness_build_trees_Earley\<^sub>L wf_\2 by blast + +end \ No newline at end of file diff --git a/thys/Earley_Parser/Limit.thy b/thys/Earley_Parser/Limit.thy new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/Limit.thy @@ -0,0 +1,146 @@ +theory Limit + imports Main +begin + +section \Slightly adjusted content from AFP/LocalLexing\ + +fun funpower :: "('a \ 'a) \ nat \ ('a \ 'a)" where + "funpower f 0 x = x" +| "funpower f (Suc n) x = f (funpower f n x)" + +definition natUnion :: "(nat \ 'a set) \ 'a set" where + "natUnion f = \ { f n | n. True }" + +definition limit :: "('a set \ 'a set) \ 'a set \ 'a set" where + "limit f x = natUnion (\ n. funpower f n x)" + +definition setmonotone :: "('a set \ 'a set) \ bool" where + "setmonotone f = (\ X. X \ f X)" + +lemma subset_setmonotone: "setmonotone f \ X \ f X" + by (simp add: setmonotone_def) + +lemma funpower_id [simp]: "funpower id n = id" + by (rule ext, induct n, simp_all) + +lemma limit_id [simp]: "limit id = id" + by (rule ext, auto simp add: limit_def natUnion_def) + +definition chain :: "(nat \ 'a set) \ bool" +where + "chain C = (\ i. C i \ C (i + 1))" + +definition continuous :: "('a set \ 'b set) \ bool" +where + "continuous f = (\ C. chain C \ (chain (f o C) \ f (natUnion C) = natUnion (f o C)))" + +lemma natUnion_upperbound: + "(\ n. f n \ G) \ (natUnion f) \ G" +by (auto simp add: natUnion_def) + +lemma funpower_upperbound: + "(\ I. I \ G \ f I \ G) \ I \ G \ funpower f n I \ G" +proof (induct n) + case 0 thus ?case by simp +next + case (Suc n) thus ?case by simp +qed + +lemma limit_upperbound: + "(\ I. I \ G \ f I \ G) \ I \ G \ limit f I \ G" +by (simp add: funpower_upperbound limit_def natUnion_upperbound) + +lemma elem_limit_simp: "x \ limit f X = (\ n. x \ funpower f n X)" +by (auto simp add: limit_def natUnion_def) + +definition pointwise :: "('a set \ 'b set) \ bool" where + "pointwise f = (\ X. f X = \ { f {x} | x. x \ X})" + +lemma natUnion_elem: "x \ f n \ x \ natUnion f" +using natUnion_def by fastforce + +lemma limit_elem: "x \ funpower f n X \ x \ limit f X" +by (simp add: limit_def natUnion_elem) + +definition pointbase :: "('a set \ 'b set) \ 'a set \ 'b set" where + "pointbase F I = \ { F X | X. finite X \ X \ I }" + +definition pointbased :: "('a set \ 'b set) \ bool" where + "pointbased f = (\ F. f = pointbase F)" + +lemma chain_implies_mono: "chain C \ mono C" +by (simp add: chain_def mono_iff_le_Suc) + +lemma setmonotone_implies_chain_funpower: + assumes setmonotone: "setmonotone f" + shows "chain (\ n. funpower f n I)" +by (simp add: chain_def setmonotone subset_setmonotone) + +lemma natUnion_subset: "(\ n. \ m. f n \ g m) \ natUnion f \ natUnion g" + by (meson natUnion_elem natUnion_upperbound subset_iff) + +lemma natUnion_eq[case_names Subset Superset]: + "(\ n. \ m. f n \ g m) \ (\ n. \ m. g n \ f m) \ natUnion f = natUnion g" +by (simp add: natUnion_subset subset_antisym) + +lemma natUnion_shift[symmetric]: + assumes chain: "chain C" + shows "natUnion C = natUnion (\ n. C (n + m))" +proof (induct rule: natUnion_eq) + case (Subset n) + show ?case using chain chain_implies_mono le_add1 mono_def by blast +next + case (Superset n) + show ?case by blast +qed + +definition regular :: "('a set \ 'a set) \ bool" +where + "regular f = (setmonotone f \ continuous f)" + +lemma regular_fixpoint: + assumes regular: "regular f" + shows "f (limit f I) = limit f I" +proof - + have setmonotone: "setmonotone f" using regular regular_def by blast + have continuous: "continuous f" using regular regular_def by blast + + let ?C = "\ n. funpower f n I" + have chain: "chain ?C" + by (simp add: setmonotone setmonotone_implies_chain_funpower) + have "f (limit f I) = f (natUnion ?C)" + using limit_def by metis + also have "f (natUnion ?C) = natUnion (f o ?C)" + by (metis continuous continuous_def chain) + also have "natUnion (f o ?C) = natUnion (\ n. f(funpower f n I))" + by (meson comp_apply) + also have "natUnion (\ n. f(funpower f n I)) = natUnion (\ n. ?C (n + 1))" + by simp + also have "natUnion (\ n. ?C(n + 1)) = natUnion ?C" + by (metis (no_types, lifting) Limit.chain_def chain natUnion_eq) + finally show ?thesis by (simp add: limit_def) +qed + +lemma fix_is_fix_of_limit: + assumes fixpoint: "f I = I" + shows "limit f I = I" +proof - + have funpower: "\ n. funpower f n I = I" + proof - + fix n :: nat + from fixpoint show "funpower f n I = I" + by (induct n, auto) + qed + show ?thesis by (simp add: limit_def funpower natUnion_def) +qed + +lemma limit_is_idempotent: "regular f \ limit f (limit f I) = limit f I" +by (simp add: fix_is_fix_of_limit regular_fixpoint) + +definition mk_regular1 :: "('b \ 'a \ bool) \ ('b \ 'a \ 'a) \ 'a set \ 'a set" where + "mk_regular1 P F I = I \ { F q x | q x. x \ I \ P q x }" + +definition mk_regular2 :: "('b \ 'a \ 'a \ bool) \ ('b \ 'a \ 'a \ 'a) \ 'a set \ 'a set" where + "mk_regular2 P F I = I \ { F q x y | q x y. x \ I \ y \ I \ P q x y }" + +end \ No newline at end of file diff --git a/thys/Earley_Parser/ROOT b/thys/Earley_Parser/ROOT new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/ROOT @@ -0,0 +1,19 @@ +chapter AFP + +session Earley_Parser (AFP) = HOL + + options [timeout = 600] + sessions + "HOL-Library" + theories + Limit + CFG + Derivations + Earley + Earley_Fixpoint + Earley_Recognizer + Earley_Parser + Examples + document_files + "root.tex" + "root.bib" + diff --git a/thys/Earley_Parser/document/root.bib b/thys/Earley_Parser/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/document/root.bib @@ -0,0 +1,65 @@ +@inproceedings{Jones:1972, +author = {Jones, C B}, +title = {Formal Development of Correct Algorithms: An Example Based on Earley's Recogniser}, +year = {1972}, +isbn = {9781450378918}, +publisher = {Association for Computing Machinery}, +address = {New York, NY, USA}, +url = {https://doi.org/10.1145/800235.807083}, +doi = {10.1145/800235.807083}, +booktitle = {Proceedings of ACM Conference on Proving Assertions about Programs}, +pages = {150–169}, +numpages = {20}, +location = {Las Cruces, New Mexico, USA} +} + +@article{Earley:1970, +author = {Earley, Jay}, +title = {An Efficient Context-Free Parsing Algorithm}, +year = {1970}, +publisher = {Association for Computing Machinery}, +address = {New York, NY, USA}, +volume = {13}, +number = {2}, +issn = {0001-0782}, +url = {https://doi.org/10.1145/362007.362035}, +doi = {10.1145/362007.362035}, +journal = {Commun. ACM}, +pages = {94–102}, +numpages = {9}, +keywords = {computational complexity, parsing, syntax analysis, context-free grammar, compilers} +} + +@article{Scott:2008, +title = {SPPF-Style Parsing From Earley Recognisers}, +journal = {Electronic Notes in Theoretical Computer Science}, +volume = {203}, +number = {2}, +pages = {53-67}, +year = {2008}, +note = {Proceedings of the Seventh Workshop on Language Descriptions, Tools, and Applications (LDTA 2007)}, +issn = {1571-0661}, +doi = {https://doi.org/10.1016/j.entcs.2008.03.044}, +url = {https://www.sciencedirect.com/science/article/pii/S1571066108001497}, +author = {Elizabeth Scott}, +keywords = {Earley parsing, cubic generalised parsing, context free languages}, +} + +@misc{Obua:2017, + title={Local Lexing}, + author={Steven Obua and Phil Scott and Jacques Fleuriot}, + year={2017}, + eprint={1702.03277}, + archivePrefix={arXiv}, + primaryClass={cs.LO} +} + +@article{LocalLexing-AFP, + author = {Steven Obua}, + title = {Local Lexing}, + journal = {Archive of Formal Proofs}, + year = {2017}, + note = {\url{https://isa-afp.org/entries/LocalLexing.html}, + Formal proof development}, + ISSN = {2150-914x}, +} diff --git a/thys/Earley_Parser/document/root.tex b/thys/Earley_Parser/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Earley_Parser/document/root.tex @@ -0,0 +1,59 @@ +\documentclass[11pt,a4paper]{article} +\usepackage[T1]{fontenc} +\usepackage{isabelle,isabellesym} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage{eurosym} + %for \ + +%\usepackage[only,bigsqcap,bigparallel,fatsemi,interleave,sslash]{stmaryrd} + %for \, \, \, \, \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + + +\begin{document} + +\title{Earley} +\author{Martin Rau} +\maketitle + +\begin{abstract} +In 1968 Earley \cite{Earley:1970} introduced his parsing algorithm capable of parsing all context-free grammars in cubic +space and time. This entry contains a formalization of an executable Earley parser. We base our development on Jones' \cite{Jones:1972} +extensive paper proof of Earley's recognizer and the formalization of context-free grammars +and derivations of Obua \cite{Obua:2017} \cite{LocalLexing-AFP}. We implement and prove correct a functional recognizer modeling Earley's +original imperative implementation and extend it with the necessary data structures to enable the construction +of parse trees following the work of Scott \cite{Scott:2008}. We then develop a functional algorithm that +builds a single parse tree and prove its correctness. Finally, we generalize this approach to an algorithm +for a complete parse forest and prove soundness. +\end{abstract} + +\tableofcontents + +% include generated text of all theories +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} \ No newline at end of file diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,756 +1,757 @@ ABY3_Protocols ADS_Functor AI_Planning_Languages_Semantics AODV AOT AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Ackermanns_not_PR Actuarial_Mathematics Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BD_Security_Compositional BNF_CC BNF_Operations BTree Balog_Szemeredi_Gowers Banach_Steinhaus Belief_Revision Bell_Numbers_Spivey BenOr_Kozen_Reif Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binary_Code_Imprimitive Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Birkhoff_Finite_Distributive_Lattices Blue_Eyes Bondy Boolean_Expression_Checkers Boolos_Curious_Inference Boolos_Curious_Inference_Automated Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CHERI-C_Memory_Model CISC-Kernel CRDT CRYSTALS-Kyber CSP_RefTK CVP_Hardness CYK CZH_Elementary_Categories CZH_Foundations CZH_Universal_Constructions CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean Clique_and_Monotone_Circuits ClockSynchInst Closest_Pair_Points CoCon CoSMeDis CoSMed CofGroups Coinductive Coinductive_Languages Collections Combinable_Wands Combinatorial_Enumeration_Algorithms Combinatorics_Words Combinatorics_Words_Graph_Lemma Combinatorics_Words_Lyndon CommCSL Commuting_Hermitian Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Bounded_Operators Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Conditional_Simplification Conditional_Transfer_Rule Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers Cook_Levin CoreC++ Core_DOM Core_SC_DOM Correctness_Algebras Cotangent_PFD_Formula Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties Crypto_Standards Cubic_Quartic_Equations DCR-ExecutionEquivalence DFS_Framework DOM_Components DPRM_Theorem DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Dedekind_Real Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Design_Theory Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic DigitsInBase Digit_Expansions Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos Distributed_Distinct_Elements Dominance_CHK DynamicArchitectures Dynamic_Tables E_Transcendental +Earley_Parser Echelon_Form EdmondsKarp_Maxflow Edwards_Elliptic_Curves_Group Efficient-Mergesort Efficient_Weighted_Path_Order Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Equivalence_Relation_Enumeration Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Eval_FO Example-Submission Executable_Randomized_Algorithms Expander_Graphs Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Axiomatic FOL_Harrison FOL_Seq_Calc1 FOL_Seq_Calc2 FOL_Seq_Calc3 FO_Theory_Rewriting FSM_Tests Factor_Algebraic_Polynomial Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF Finite_Fields Finitely_Generated_Abelian_Groups First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Fishers_Inequality Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Foundation_of_geometry Fourier Free-Boolean-Algebra Free-Groups Frequency_Moments Fresh_Identifiers FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC GaleStewart_Games Gale_Shapley Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic Given_Clause_Loops GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Gray_Codes Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Grothendieck_Schemes Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Hahn_Jordan_Decomposition Hales_Jewett Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms HoareForDivergence Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL HyperHoareLogic Hyperdual IEEE_Floating_Point IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IMP_Compiler IMP_Compiler_Reuse IP_Addresses Imperative_Insertion_Sort Implicational_Logic Impossible_Geometry Incompleteness Incredible_Proof_Machine Independence_CH Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpolation_Polynomials_HOL_Algebra Interpreter_Optimizations Interval_Arithmetic_Word32 Intro_Dest_Elim Involutions2Squares Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Irrationals_From_THEBOOK IsaGeoCoq IsaNet Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Khovanskii_Theorem Kleene_Algebra Kneser_Cauchy_Davenport Knights_Tour Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LP_Duality LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option Lifting_the_Exponent LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Logging_Independent_Anonymity Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MDP-Algorithms MDP-Rewards MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Maximum_Segment_Sum Median_Method Median_Of_Medians_Selection Menger Mereology Mersenne_Primes Metalogic_ProofChecker MHComputation MiniML MiniSail Minimal_SSA Minkowskis_Theorem Minsky_Machines MLSS_Decision_Proc Modal_Logics_for_NTS Modular_Assembly_Kit_Security Modular_arithmetic_LLL_and_HNF_algorithms Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Multirelations_Heterogeneous Multiset_Ordering_NPC Multitape_To_Singletape_TM Myhill-Nerode Name_Carrying_Type_Inference Nano_JSON Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers No_FTL_observers_Gen_Rel Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Number_Theoretic_Transform Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PAL PAPP_Impossibility PCF PLM POPLmark-deBruijn PSemigroupsConvolution Package_logic Padic_Field Padic_Ints Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Pluennecke_Ruzsa_Inequality Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Prefix_Free_Code_Combinators Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Probability_Inequality_Completeness Program-Conflict-Analysis Progress_Tracking Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Logic_Class Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem Public_Announcement_Logic QHLProver QR_Decomposition Quantales Quantifier_Elimination_Hybrid Quasi_Borel_Spaces Quaternions Query_Optimization Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Real_Power Real_Time_Deque Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Registers Regression_Test_Selection Regular-Sets Regular_Algebras Regular_Tree_Relations Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rensets Rep_Fin_Groups ResiduatedTransitionSystem Residuated_Lattices Resolution_FOL Rewrite_Properties_Reduction Rewriting_Z Ribbon_Proofs Risk_Free_Lending Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Roth_Arithmetic_Progressions Routing Roy_Floyd_Warshall SATSolverVerification SCC_Bloemen_Sequential SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Safe_Range_RC Saturation_Framework Saturation_Framework_Extensions Sauer_Shelah_Lemma Schutz_Spacetime Schwartz_Zippel Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL Separation_Logic_Unbounded SequentInvertibility Shadow_DOM Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Clause_Learning Simple_Firewall Simplex Simplicial_complexes_and_boolean_functions SimplifiedOntologicalArgument Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Solidity Sophomores_Dream Sort_Encodings Source_Coding_Theorem SpecCheck Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Stalnaker_Logic Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code StrictOmegaCategories Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Suppes_Theorem Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Synthetic_Completeness Szemeredi_Regularity Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Three_Circles Three_Squares Timed_Automata Topological_Semantics Topology TortoiseHare TsirelsonBound Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Transitive_Models Treaps Tree-Automata Tree_Decomposition Tree_Enumeration Triangle Trie Turans_Graph_Theorem Twelvefold_Way Two_Generated_Word_Monoids_Intersection Tycon Types_Tableaus_and_Goedels_God Types_To_Sets_Extension UPF UPF_Firewall UTP Undirected_Graph_Theory Universal_Hash_Families Universal_Turing_Machine UpDown_Scheme VYDRA_MDL Valuation Van_Emde_Boas_Trees Van_der_Waerden VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves Virtual_Substitution VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Weighted_Arithmetic_Geometric_Mean Weighted_Path_Order Well_Quasi_Orders Wetzels_Problem Winding_Number_Eval Word_Lib WorkerWrapper X86_Semantics XML Youngs_Inequality ZFC_in_HOL Zeckendorf Zeta_3_Irrational Zeta_Function pGCL