diff --git a/thys/Auto2_HOL/HOL/Arith_Thms.thy b/thys/Auto2_HOL/HOL/Arith_Thms.thy --- a/thys/Auto2_HOL/HOL/Arith_Thms.thy +++ b/thys/Auto2_HOL/HOL/Arith_Thms.thy @@ -1,252 +1,252 @@ (* File: Arith_Thms.thy Author: Bohua Zhan Setup for proof steps related to arithmetic, mostly on natural numbers. *) section \Setup for arithmetic\ theory Arith_Thms imports Order_Thms HOL.Binomial begin (* Reducing inequality on natural numbers. *) theorem reduce_le_plus_consts: "(x::nat) + n1 \ y + n2 \ x \ y + (n2-n1)" by simp theorem reduce_le_plus_consts': "n1 \ n2 \ (x::nat) + n1 \ y + n2 \ x + (n1-n2) \ y" by simp theorem reduce_less_plus_consts: "(x::nat) + n1 < y + n2 \ x < y + (n2-n1)" by simp theorem reduce_less_plus_consts': "n1 \ n2 \ (x::nat) + n1 < y + n2 \ x + (n1-n2) < y" by simp (* To normal form. *) theorem norm_less_lminus: "(x::nat) - n < y \ x \ y + (n-1)" by simp theorem norm_less_lplus: "(x::nat) + n < y \ x + (n+1) \ y" by simp theorem norm_less_rminus: "(x::nat) < y - n \ x + (n+1) \ y" by simp theorem norm_less_rplus: "(x::nat) < y + n \ x \ y + (n-1)" by simp theorem norm_less: "(x::nat) < y \ x + 1 \ y" by simp theorem norm_le_lminus: "(x::nat) - n \ y \ x \ y + n" by simp theorem norm_le_rminus: "(x::nat) \ y - n \ x \ y + 0" by simp theorem norm_le: "(x::nat) \ y \ x \ y + 0" by simp theorem norm_le_lplus0: "(x::nat) + 0 \ y \ x \ y + 0" by simp (* Transitive resolve. *) theorem trans_resolve1: "n1 > 0 \ (x::nat) + n1 \ y \ (y::nat) + n2 \ x \ False" by simp theorem trans_resolve2: "n1 > n2 \ (x::nat) + n1 \ y \ (y::nat) \ x + n2 \ False" by simp (* Transitive. *) theorem trans1: "(x::nat) + n1 \ y \ y + n2 \ z \ x + (n1+n2) \ z" by simp theorem trans2: "(x::nat) \ y + n1 \ y \ z + n2 \ x \ z + (n1+n2)" by simp theorem trans3: "(x::nat) + n1 \ y \ y \ z + n2 \ x \ z + (n2-n1)" by simp theorem trans4: "n1 > n2 \ (x::nat) + n1 \ y \ y \ z + n2 \ x + (n1-n2) \ z" by simp theorem trans5: "(x::nat) \ y + n1 \ y + n2 \ z \ x \ z + (n1-n2)" by simp theorem trans6: "n2 > n1 \ (x::nat) \ y + n1 \ y + n2 \ z \ x + (n2-n1) \ z" by simp (* Resolve. *) theorem single_resolve: "n > 0 \ (x::nat) + n \ x \ False" by simp theorem single_resolve_const: "n > 0 \ (x::nat) + n \ 0 \ False" by simp (* Comparison with constants. *) theorem cv_const1: "(x::nat) + n \ y \ 0 + (x+n) \ y" by simp (* x is const *) theorem cv_const2: "(x::nat) + n \ y \ x \ 0 + (y-n)" by simp (* y is const *) theorem cv_const3: "y < n \ (x::nat) + n \ y \ x + (n-y) \ 0" by simp (* y is const (contradiction with 0 \ x) *) theorem cv_const4: "(x::nat) \ y + n \ 0 + (x-n) \ y" by simp (* x is const *) theorem cv_const5: "(x::nat) \ y + n \ 0 \ y + (n-x)" by simp (* x is const (trivial) *) theorem cv_const6: "(x::nat) \ y + n \ x \ 0 + (y+n)" by simp (* y is const *) (* Misc *) theorem nat_eq_to_ineqs: "(x::nat) = y + n \ x \ y + n \ x \ y + n" by simp theorem nat_ineq_impl_not_eq: "(x::nat) + n \ y \ n > 0 \ x \ y" by simp theorem eq_to_ineqs: "(x::nat) \ y \ x \ y \ y \ x" by simp theorem ineq_to_eqs1: "(x::nat) \ y + 0 \ y \ x + 0 \ x = y" by simp ML_file \arith.ML\ ML_file \order.ML\ ML_file \order_test.ML\ setup \register_wellform_data ("(a::nat) - b", ["a \ b"])\ setup \add_prfstep_check_req ("(a::nat) - b", "(a::nat) \ b")\ (* Normalize any expression to "a - b" form. *) lemma nat_sub_norm: "(a::nat) = a - 0 \ a \ 0" by simp (* Adding and subtracting two normalized expressions. *) lemma nat_sub1: "(a::nat) \ b \ c \ d \ (a - b) + (c - d) = (a + c) - (b + d) \ a + c \ b + d" by simp lemma nat_sub2: "(a::nat) \ b \ c \ d \ a - b \ c - d \ (a - b) - (c - d) = (a + d) - (b + c) \ a + d \ b + c" by simp lemma nat_sub3: "(a::nat) \ b \ c \ d \ (a - b) * (c - d) = (a * c + b * d) - (a * d + b * c) \ a * c + b * d \ a * d + b * c" by (smt diff_mult_distrib mult.commute mult_le_mono1 nat_sub2) (* Cancel identical terms on two sides, yielding a normalized expression. *) lemma nat_sub_combine: "(a::nat) + b \ c + b \ (a + b) - (c + b) = a - c \ a \ c" by simp lemma nat_sub_combine2: "n \ m \ (a::nat) + b * n \ c + b * m \ (a + b * n) - (c + b * m) = (a + b * (n - m)) - c \ a + b * (n - m) \ c \ n \ m" by (simp add: add.commute right_diff_distrib') lemma nat_sub_combine3: "n \ m \ (a::nat) + b * n \ c + b * m \ (a + b * n) - (c + b * m) = a - (c + b * (m - n)) \ a \ c + b * (m - n) \ m \ n" by (smt add.commute mult.commute nat_diff_add_eq2 nat_le_add_iff1) ML_file \nat_sub.ML\ ML_file \nat_sub_test.ML\ (* Ordering on Nats. *) lemma le_neq_implies_less' [forward]: "(m::nat) \ n \ m \ n \ m < n" by simp lemma le_zero_to_equal_zero [forward]: "(n::nat) \ 0 \ n = 0" by simp lemma less_one_to_equal_zero [forward]: "(n::nat) < 1 \ n = 0" by simp setup \add_backward_prfstep_cond @{thm Nat.mult_le_mono1} [with_cond "?k \ 1"]\ setup \add_resolve_prfstep @{thm Nat.not_add_less1}\ lemma not_minus_less [resolve]: "\(i::nat) < (i - j)" by simp lemma nat_le_prod_with_same [backward]: "m \ 0 \ (n::nat) \ m * n" by simp lemma nat_le_prod_with_le [backward1]: "k \ 0 \ (n::nat) \ m \ (n::nat) \ k * m" using le_trans nat_le_prod_with_same by blast lemma nat_plus_le_to_less [backward1]: "b \ 0 \ (a::nat) + b \ c \ a < c" by simp lemma nat_plus_le_to_less2 [backward1]: "a \ 0 \ (a::nat) + b \ c \ b < c" by simp setup \add_forward_prfstep @{thm add_right_imp_eq}\ setup \add_forward_prfstep @{thm add_left_imp_eq}\ -setup \add_rewrite_rule_cond @{thm Nat.le_diff_conv2} [with_term "?i + ?k"]\ +setup \add_rewrite_rule_cond @{thm le_diff_conv2} [with_term "?i + ?k"]\ lemma nat_less_diff_conv: "(i::nat) < j - k \ i + k < j" by simp setup \add_forward_prfstep_cond @{thm nat_less_diff_conv} [with_cond "?k \ ?NUMC", with_term "?i + ?k"]\ lemma Nat_le_diff_conv2_same [forward]: "i \ j \ (i::nat) \ i - j \ j = 0" by simp lemma nat_gt_zero [forward]: "b - a > 0 \ b > (a::nat)" by simp lemma n_minus_1_less_n: "(n::nat) \ 1 \ n - 1 < n" by simp setup \add_forward_prfstep_cond @{thm n_minus_1_less_n} [with_term "?n - 1"]\ (* Monotonicity of ordering *) setup \add_backward_prfstep @{thm Nat.diff_le_mono}\ setup \add_backward2_prfstep @{thm Nat.diff_less_mono}\ setup \add_backward_prfstep @{thm Nat.mult_le_mono2}\ setup \add_resolve_prfstep @{thm Nat.le_add1}\ setup \add_resolve_prfstep @{thm Nat.le_add2}\ setup \add_backward_prfstep @{thm add_left_mono}\ setup \add_backward_prfstep @{thm add_right_mono}\ lemma add_mono_neutr [backward]: "(0::'a::linordered_ring) \ b \ a \ a + b" by simp lemma add_mono_neutl [backward]: "(0::'a::linordered_ring) \ b \ a \ b + a" by simp setup \add_forward_prfstep @{thm add_less_imp_less_left}\ lemma sum_le_zero1 [forward]: "(a::'a::linordered_ring) + b < 0 \ a \ 0 \ b < 0" by (meson add_less_same_cancel1 less_le_trans) lemma less_sum1 [backward]: "b > 0 \ a < a + (b::nat)" by auto setup \add_backward_prfstep @{thm Nat.trans_less_add2}\ setup \add_backward_prfstep @{thm Nat.add_less_mono1}\ setup \add_backward1_prfstep @{thm Nat.add_less_mono}\ setup \add_backward1_prfstep @{thm Nat.add_le_mono}\ setup \add_backward1_prfstep @{thm add_increasing2}\ setup \add_backward1_prfstep @{thm add_mono}\ setup \add_backward_prfstep @{thm add_strict_left_mono}\ setup \add_backward1_prfstep @{thm Nat.mult_le_mono}\ (* Addition. *) theorem nat_add_eq_self_zero [forward]: "(m::nat) = m + n \ n = 0" by simp theorem nat_add_eq_self_zero' [forward]: "(m::nat) = n + m \ n = 0" by simp theorem nat_mult_2: "(a::nat) + a = 2 * a" by simp setup \add_rewrite_rule_cond @{thm nat_mult_2} [with_cond "?a \ 0"]\ theorem plus_one_non_zero [resolve]: "\(n::nat) + 1 = 0" by simp (* Diff. *) lemma nat_same_minus_ge [forward]: "(c::nat) - a \ c - b \ a \ c \ a \ b" by arith lemma diff_eq_zero [forward]: "(k::nat) \ j \ j - k = 0 \ j = k" by simp lemma diff_eq_zero' [forward]: "(k::nat) \ j \ j - k + i = j \ k = i" by simp (* Divides. *) theorem dvd_defD1 [resolve]: "(a::nat) dvd b \ \k. b = a * k" using dvdE by blast theorem dvd_defD2 [resolve]: "(a::nat) dvd b \ \k. b = k * a" by (metis dvd_mult_div_cancel mult.commute) setup \add_forward_prfstep @{thm Nat.dvd_imp_le}\ theorem dvd_ineq2 [forward]: "(k::nat) dvd n \ n > 0 \ k \ 1" by (simp add: Suc_leI dvd_pos_nat) setup \add_forward_prfstep_cond @{thm dvd_trans} (with_conds ["?a \ ?b", "?b \ ?c", "?a \ ?c"])\ setup \add_forward_prfstep_cond @{thm Nat.dvd_antisym} [with_cond "?m \ ?n"]\ theorem dvd_cancel [backward1]: "c > 0 \ (a::nat) * c dvd b * c \ a dvd b" by simp setup \add_forward_prfstep (equiv_forward_th @{thm dvd_add_right_iff})\ (* Divisibility: existence. *) setup \add_resolve_prfstep @{thm dvd_refl}\ theorem exists_n_dvd_n [backward]: "P (n::nat) \ \k. k dvd n \ P k" using dvd_refl by blast setup \add_resolve_prfstep @{thm one_dvd}\ theorem any_n_dvd_0 [forward]: "\ (\ k. k dvd (0::nat) \ P k) \ \ (\ k. P k)" by simp theorem n_dvd_one: "(n::nat) dvd 1 \ n = 1" by simp setup \add_forward_prfstep_cond @{thm n_dvd_one} [with_cond "?n \ 1"]\ (* Products. *) setup \add_rewrite_rule @{thm mult_zero_left}\ lemma prod_ineqs1 [forward]: "(m::nat) * k > 0 \ m > 0 \ k > 0" by simp lemma prod_ineqs2 [backward]: "(k::nat) > 0 \ m \ m * k" by simp theorem prod_cancel: "(a::nat) * b = a * c \ a > 0 \ b = c" by auto setup \add_forward_prfstep_cond @{thm prod_cancel} [with_cond "?b \ ?c"]\ theorem mult_n1n: "(n::nat) = m * n \ n > 0 \ m = 1" by auto setup \add_forward_prfstep_cond @{thm mult_n1n} [with_cond "?m \ 1"]\ theorem prod_is_one [forward]: "(x::nat) * y = 1 \ x = 1" by simp theorem prod_dvd_intro [backward]: "(k::nat) dvd m \ k dvd n \ k dvd m * n" using dvd_mult dvd_mult2 by blast (* Definition of gcd. *) setup \add_forward_prfstep_cond @{thm gcd_dvd1} [with_term "gcd ?a ?b"]\ setup \add_forward_prfstep_cond @{thm gcd_dvd2} [with_term "gcd ?a ?b"]\ (* Coprimality. *) setup \add_rewrite_rule_bidir @{thm coprime_iff_gcd_eq_1}\ lemma coprime_exp [backward]: "coprime d a \ coprime (d::nat) (a ^ n)" by simp setup \add_backward_prfstep @{thm coprime_exp}\ setup \add_rewrite_rule @{thm gcd.commute}\ lemma coprime_dvd_mult [backward1]: "coprime (a::nat) b \ a dvd c * b \ a dvd c" by (metis coprime_dvd_mult_left_iff) lemma coprime_dvd_mult' [backward1]: "coprime (a::nat) b \ a dvd b * c \ a dvd c" by (metis coprime_dvd_mult_right_iff) theorem coprime_dvd [forward]: "coprime (a::nat) b \ p dvd a \ p > 1 \ \ p dvd b" using coprime_common_divisor_nat by blast (* Powers. *) setup \add_rewrite_rule @{thm power_0}\ theorem power_ge_0 [rewrite]: "m \ 0 \ p ^ m = p * (p ^ (m - 1))" by (simp add: power_eq_if) setup \add_rewrite_rule_cond @{thm power_one} [with_cond "?n \ 0"]\ setup \add_rewrite_rule_cond @{thm power_one_right} [with_cond "?a \ 1"]\ setup \add_gen_prfstep ("power_case_intro", [WithTerm @{term_pat "?p ^ (?FREE::nat)"}, CreateCase @{term_pat "(?FREE::nat) = 0"}])\ lemma one_is_power_of_any [resolve]: "\i. (1::nat) = a ^ i" by (metis power.simps(1)) setup \add_rewrite_rule @{thm power_Suc}\ theorem power_dvd [forward]: "(p::nat)^n dvd a \ n \ 0 \ p dvd a" using dvd_power dvd_trans by blast theorem power_eq_one: "(b::nat) ^ n = 1 \ b = 1 \ n = 0" by (metis One_nat_def Suc_lessI nat_zero_less_power_iff power_0 power_inject_exp) setup \add_forward_prfstep_cond @{thm power_eq_one} (with_conds ["?b \ 1", "?n \ 0"])\ (* Factorial. *) theorem fact_ge_1_nat: "fact (n::nat) \ (1::nat)" by simp setup \add_forward_prfstep_cond @{thm fact_ge_1_nat} [with_term "fact ?n"]\ setup \add_backward1_prfstep @{thm dvd_fact}\ (* Successor function. *) setup \add_rewrite_rule @{thm Nat.Suc_eq_plus1}\ setup \add_backward_prfstep @{thm Nat.gr0_implies_Suc}\ (* Cases *) setup \fold add_rewrite_rule @{thms Nat.nat.case}\ (* Induction. *) lemma nat_cases: "P 0 \ (\n. P (Suc n)) \ P n" using nat_induct by auto (* div *) declare times_div_less_eq_dividend [resolve] setup \ add_var_induct_rule @{thm nat_induct} #> add_strong_induct_rule @{thm nat_less_induct} #> add_cases_rule @{thm nat_cases} \ end diff --git a/thys/Gauss_Jordan/Rank.thy b/thys/Gauss_Jordan/Rank.thy --- a/thys/Gauss_Jordan/Rank.thy +++ b/thys/Gauss_Jordan/Rank.thy @@ -1,46 +1,46 @@ (* Title: Rank.thy Author: Jose Divasón Author: Jesús Aransay Maintainer: Jose Divasón *) section\Rank of a matrix\ theory Rank imports Rank_Nullity_Theorem.Dim_Formula begin subsection\Row rank, column rank and rank\ text\Definitions of row rank, column rank and rank\ definition row_rank :: "'a::{field}^'n^'m=>nat" where "row_rank A = vec.dim (row_space A)" definition col_rank :: "'a::{field}^'n^'m=>nat" where "col_rank A = vec.dim (col_space A)" lemma rank_def: "rank A = row_rank A" by (auto simp: row_rank_def row_rank_def_gen row_space_def) subsection\Properties\ lemma rrk_is_preserved: fixes A::"'a::{field}^'cols^'rows::{finite, wellorder}" and P::"'a::{field}^'rows::{finite, wellorder}^'rows::{finite, wellorder}" assumes inv_P: "invertible P" shows "row_rank A = row_rank (P**A)" by (metis row_space_is_preserved row_rank_def inv_P) lemma crk_is_preserved: fixes A::"'a::{field}^'cols::{finite, wellorder}^'rows" and P::"'a::{field}^'rows^'rows" assumes inv_P: "invertible P" shows "col_rank A = col_rank (P**A)" using rank_nullity_theorem_matrices unfolding ncols_def - by (metis col_rank_def inv_P nat_add_left_cancel null_space_is_preserved) + by (metis col_rank_def inv_P add_left_cancel null_space_is_preserved) end diff --git a/thys/List-Infinite/CommonArith/Util_Nat.thy b/thys/List-Infinite/CommonArith/Util_Nat.thy --- a/thys/List-Infinite/CommonArith/Util_Nat.thy +++ b/thys/List-Infinite/CommonArith/Util_Nat.thy @@ -1,311 +1,311 @@ (* Title: Util_Nat.thy Date: Oct 2006 Author: David Trachtenherz *) section \Results for natural arithmetics\ theory Util_Nat imports Main begin subsection \Some convenience arithmetic lemmata\ lemma add_1_Suc_conv: "m + 1 = Suc m" by simp lemma sub_Suc0_sub_Suc_conv: "b - a - Suc 0 = b - Suc a" by simp lemma Suc_diff_Suc: "m < n \ Suc (n - Suc m) = n - m" apply (rule subst[OF sub_Suc0_sub_Suc_conv]) apply (rule Suc_pred) apply (simp only: zero_less_diff) done lemma nat_grSuc0_conv: "(Suc 0 < n) = (n \ 0 \ n \ Suc 0)" by fastforce lemma nat_geSucSuc0_conv: "(Suc (Suc 0) \ n) = (n \ 0 \ n \ Suc 0)" by fastforce lemma nat_lessSucSuc0_conv: "(n < Suc (Suc 0)) = (n = 0 \ n = Suc 0)" by fastforce lemma nat_leSuc0_conv: "(n \ Suc 0) = (n = 0 \ n = Suc 0)" by fastforce lemma mult_pred: "(m - Suc 0) * n = m * n - n" by (simp add: diff_mult_distrib) lemma mult_pred_right: "m * (n - Suc 0) = m * n - m" by (simp add: diff_mult_distrib2) lemma gr_implies_gr0: "m < (n::nat) \ 0 < n" by simp corollary mult_cancel1_gr0: " (0::nat) < k \ (k * m = k * n) = (m = n)" by simp corollary mult_cancel2_gr0: " (0::nat) < k \ (m * k = n * k) = (m = n)" by simp corollary mult_le_cancel1_gr0: " (0::nat) < k \ (k * m \ k * n) = (m \ n)" by simp corollary mult_le_cancel2_gr0: " (0::nat) < k \ (m * k \ n * k) = (m \ n)" by simp lemma gr0_imp_self_le_mult1: "0 < (k::nat) \ m \ m * k" by (drule Suc_leI, drule mult_le_mono[OF order_refl], simp) lemma gr0_imp_self_le_mult2: "0 < (k::nat) \ m \ k * m" by (subst mult.commute, rule gr0_imp_self_le_mult1) lemma less_imp_Suc_mult_le: "m < n \ Suc m * k \ n * k" by (rule mult_le_mono1, simp) lemma less_imp_Suc_mult_pred_less: "\ m < n; 0 < k \ \ Suc m * k - Suc 0 < n * k" apply (rule Suc_le_lessD) apply (simp only: Suc_pred[OF nat_0_less_mult_iff[THEN iffD2, OF conjI, OF zero_less_Suc]]) apply (rule less_imp_Suc_mult_le, assumption) done lemma ord_zero_less_diff: "(0 < (b::'a::ordered_ab_group_add) - a) = (a < b)" by (simp add: less_diff_eq) lemma ord_zero_le_diff: "(0 \ (b::'a::ordered_ab_group_add) - a) = (a \ b)" by (simp add: le_diff_eq) text \\diff_diff_right\ in rule format\ lemmas diff_diff_right = Nat.diff_diff_right[rule_format] lemma less_add1: "(0::nat) < j \ i < i + j" by simp lemma less_add2: "(0::nat) < j \ i < j + i" by simp lemma add_lessD2: "i + j < (k::nat) \ j < k" by simp lemma add_le_mono2: "i \ (j::nat) \ k + i \ k + j" by simp lemma add_less_mono2: "i < (j::nat) \ k + i < k + j" by simp lemma diff_less_self: "\ (0::nat) < i; 0 < j \ \ i - j < i" by simp lemma ge_less_neq_conv: "((a::'a::linorder) \ n) = (\x. x < a \ n \ x)" and le_greater_neq_conv: "(n \ (a::'a::linorder)) = (\x. a < x \ n \ x)" by (subst linorder_not_less[symmetric], blast)+ lemma greater_le_neq_conv: "((a::'a::linorder) < n) = (\x. x \ a \ n \ x)" and less_ge_neq_conv: "(n < (a::'a::linorder)) = (\x. a \ x \ n \ x)" by (subst linorder_not_le[symmetric], blast)+ text \Lemmas for @term{abs} function\ lemma leq_pos_imp_abs_leq: "\ 0 \ (a::'a::ordered_ab_group_add_abs); a \ b \ \ \a\ \ \b\" by simp lemma leq_neg_imp_abs_geq: "\ (a::'a::ordered_ab_group_add_abs) \ 0; b \ a \ \ \a\ \ \b\" by simp lemma abs_range: "\ 0 \ (a::'a::{ordered_ab_group_add_abs,abs_if}); -a \ x; x \ a \ \ \x\ \ a" apply (clarsimp simp: abs_if) apply (rule neg_le_iff_le[THEN iffD1], simp) done text \Lemmas for @term{sgn} function\ lemma sgn_abs:"(x::'a::linordered_idom) \ 0 \ \sgn x\ = 1" by (case_tac "x < 0", simp+) lemma sgn_mult_abs:"\x\ * \sgn (a::'a::linordered_idom)\ = \x * sgn a\" by (fastforce simp add: sgn_if abs_if) lemma abs_imp_sgn_abs: "\a\ = \b\ \ \sgn (a::'a::linordered_idom)\ = \sgn b\" by (fastforce simp add: abs_if) lemma sgn_mono: "a \ b \ sgn (a::'a::{linordered_idom,linordered_semidom}) \ sgn b" by (auto simp add: sgn_if) subsection \Additional facts about inequalities\ lemma add_diff_le: "k \ n \ m + k - n \ (m::nat)" by (case_tac "m + k < n", simp_all) lemma less_add_diff: "k < (n::nat) \ m < n + m - k" by (rule add_less_imp_less_right[of _ k], simp) lemma add_diff_less: "\ k < n; 0 < m \ \ m + k - n < (m::nat)" by (case_tac "m + k < n", simp_all) lemma add_le_imp_le_diff1: "i + k \ j \ i \ j - (k::nat)" by (case_tac "k \ j", simp_all) lemma add_le_imp_le_diff2: "k + i \ j \ i \ j - (k::nat)" by simp lemma diff_less_imp_less_add: "j - (k::nat) < i \ j < i + k" by simp lemma diff_less_conv: "0 < i \ (j - (k::nat) < i) = (j < i + k)" by (safe, simp_all) lemma le_diff_swap: "\ i \ (k::nat); j \ k \ \ (k - j \ i) = (k - i \ j)" by (safe, simp_all) lemma diff_less_imp_swap: "\ 0 < (i::nat); k - i < j \ \ (k - j < i)" by simp lemma diff_less_swap: "\ 0 < (i::nat); 0 < j \ \ (k - j < i) = (k - i < j)" by (blast intro: diff_less_imp_swap) lemma less_diff_imp_less: "(i::nat) < j - m \ i < j" by simp lemma le_diff_imp_le: "(i::nat) \ j - m \ i \ j" by simp lemma less_diff_le_imp_less: "\ (i::nat) < j - m; n \ m \ \ i < j - n" by simp lemma le_diff_le_imp_le: "\ (i::nat) \ j - m; n \ m \ \ i \ j - n" by simp lemma le_imp_diff_le: "(j::nat) \ k \ j - n \ k" by simp subsection \Inequalities for Suc and pred\ corollary less_eq_le_pred: "0 < (n::nat) \ (m < n) = (m \ n - Suc 0)" by (safe, simp_all) corollary less_imp_le_pred: "m < n \ m \ n - Suc 0" by simp corollary le_pred_imp_less: "\ 0 < n; m \ n - Suc 0 \ \ m < n" by simp corollary pred_less_eq_le: "0 < m \ (m - Suc 0 < n) = (m \ n)" by (safe, simp_all) corollary pred_less_imp_le: "m - Suc 0 < n \ m \ n" by simp corollary le_imp_pred_less: "\ 0 < m; m \ n \ \ m - Suc 0 < n" by simp lemma diff_add_inverse_Suc: "n < m \ n + (m - Suc n) = m - Suc 0" by simp lemma pred_mono: "\ m < n; 0 < m \ \ m - Suc 0 < n - Suc 0" by simp corollary pred_Suc_mono: "\ m < Suc n; 0 < m \ \ m - Suc 0 < n" by simp lemma Suc_less_pred_conv: "(Suc m < n) = (m < n - Suc 0)" by (safe, simp_all) lemma Suc_le_pred_conv: "0 < n \ (Suc m \ n) = (m \ n - Suc 0)" by (safe, simp_all) lemma Suc_le_imp_le_pred: "Suc m \ n \ m \ n - Suc 0" by simp subsection \Additional facts about cancellation in (in-)equalities\ lemma diff_cancel_imp_eq: "\ 0 < (n::nat); n + i - j = n \ \ i = j" by simp lemma nat_diff_left_cancel_less: "k - m < k - (n::nat) \ n < m" by simp lemma nat_diff_right_cancel_less: "n - k < (m::nat) - k \ n < m" by simp lemma nat_diff_left_cancel_le1: "\ k - m \ k - (n::nat); m < k \ \ n \ m" by simp lemma nat_diff_left_cancel_le2: "\ k - m \ k - (n::nat); n \ k \ \ n \ m" by simp lemma nat_diff_right_cancel_le1: "\ m - k \ n - (k::nat); k < m \ \ m \ n" by simp lemma nat_diff_right_cancel_le2: "\ m - k \ n - (k::nat); k \ n \ \ m \ n" by simp lemma nat_diff_left_cancel_eq1: "\ k - m = k - (n::nat); m < k \ \ m = n" by simp lemma nat_diff_left_cancel_eq2: "\ k - m = k - (n::nat); n < k \ \ m = n" by simp lemma nat_diff_right_cancel_eq1: "\ m - k = n - (k::nat); k < m \ \ m = n" by simp lemma nat_diff_right_cancel_eq2: "\ m - k = n - (k::nat); k < n \ \ m = n" by simp lemma eq_diff_left_iff: "\ (m::nat) \ k; n \ k\ \ (k - m = k - n) = (m = n)" by (safe, simp_all) lemma eq_imp_diff_eq: "m = (n::nat) \ m - k = n - k" by simp text \List of definitions and lemmas\ thm Nat.add_Suc_right add_1_Suc_conv sub_Suc0_sub_Suc_conv thm Nat.mult_cancel1 Nat.mult_cancel2 mult_cancel1_gr0 mult_cancel2_gr0 thm Nat.add_lessD1 add_lessD2 thm Nat.zero_less_diff ord_zero_less_diff ord_zero_le_diff thm Nat.le_add_diff add_diff_le less_add_diff add_diff_less thm - Nat.le_diff_conv Nat.le_diff_conv2 + Nat.le_diff_conv le_diff_conv2 Nat.less_diff_conv diff_less_imp_less_add diff_less_conv thm le_diff_swap diff_less_imp_swap diff_less_swap thm less_diff_imp_less le_diff_imp_le thm less_diff_le_imp_less le_diff_le_imp_le thm Nat.less_imp_diff_less le_imp_diff_le thm Nat.less_Suc_eq_le less_eq_le_pred less_imp_le_pred le_pred_imp_less thm Nat.Suc_le_eq pred_less_eq_le pred_less_imp_le le_imp_pred_less thm diff_cancel_imp_eq thm diff_add_inverse_Suc thm Nat.nat_add_left_cancel_less Nat.nat_add_left_cancel_le Nat.nat_add_right_cancel - Nat.nat_add_left_cancel + Nat.add_left_cancel Nat.eq_diff_iff Nat.less_diff_iff Nat.le_diff_iff thm nat_diff_left_cancel_less nat_diff_right_cancel_less thm nat_diff_left_cancel_le1 nat_diff_left_cancel_le2 nat_diff_right_cancel_le1 nat_diff_right_cancel_le2 thm nat_diff_left_cancel_eq1 nat_diff_left_cancel_eq2 nat_diff_right_cancel_eq1 nat_diff_right_cancel_eq2 thm Nat.eq_diff_iff eq_diff_left_iff thm - Nat.nat_add_right_cancel Nat.nat_add_left_cancel + Nat.nat_add_right_cancel Nat.add_left_cancel Nat.diff_le_mono eq_imp_diff_eq end diff --git a/thys/List-Infinite/ListInf/List2.thy b/thys/List-Infinite/ListInf/List2.thy --- a/thys/List-Infinite/ListInf/List2.thy +++ b/thys/List-Infinite/ListInf/List2.thy @@ -1,1194 +1,1194 @@ (* Title: List2.thy Date: Oct 2006 Author: David Trachtenherz *) section \Additional definitions and results for lists\ theory List2 imports "../CommonSet/SetIntervalCut" begin subsection \Additional definitions and results for lists\ text \ Infix syntactical abbreviations for operators @{term take} and @{term drop}. The abbreviations resemble to the operator symbols used later for take and drop operators on infinite lists in ListInf.\ abbreviation f_take' :: "'a list \ nat \ 'a list" (infixl "\" 100) where "xs \ n \ take n xs" abbreviation f_drop' :: "'a list \ nat \ 'a list" (infixl "\" 100) where "xs \ n \ drop n xs" lemma append_eq_Cons: "[x] @ xs = x # xs" by simp lemma length_Cons: "length (x # xs) = Suc (length xs)" by simp lemma length_snoc: "length (xs @ [x]) = Suc (length xs)" by simp subsubsection \Additional lemmata about list emptiness\ lemma length_greater_imp_not_empty:"n < length xs \ xs \ []" by fastforce lemma length_ge_Suc_imp_not_empty:"Suc n \ length xs \ xs \ []" by fastforce lemma length_take_le: "length (xs \ n) \ length xs" by simp lemma take_not_empty_conv:"(xs \ n \ []) = (0 < n \ xs \ [])" by simp lemma drop_not_empty_conv:"(xs \ n \ []) = (n < length xs)" by fastforce lemma zip_eq_Nil: "(zip xs ys = []) = (xs = [] \ ys = [])" by (force simp: length_0_conv[symmetric] min_def simp del: length_0_conv) lemma zip_not_empty_conv: "(zip xs ys \ []) = (xs \ [] \ ys \ [])" by (simp add: zip_eq_Nil) subsubsection \Additional lemmata about @{term take}, @{term drop}, @{term hd}, @{term last}, \nth\ and \filter\\ lemma nth_tl_eq_nth_Suc: " Suc n \ length xs \ (tl xs) ! n = xs ! Suc n" by (rule hd_Cons_tl[OF length_ge_Suc_imp_not_empty, THEN subst], simp+) corollary nth_tl_eq_nth_Suc2: " n < length xs \ (tl xs) ! n = xs ! Suc n" by (simp add: nth_tl_eq_nth_Suc) lemma hd_eq_first: "xs \ [] \ xs ! 0 = hd xs" by (induct xs, simp_all) corollary take_first:"xs \ [] \ xs \ (Suc 0) = [xs ! 0]" by (induct xs, simp_all) corollary take_hd:"xs \ [] \ xs \ (Suc 0) = [hd xs]" by (simp add: take_first hd_eq_first) theorem last_nth: "xs \ [] \ last xs = xs ! (length xs - Suc 0)" by (simp add: last_conv_nth) lemma last_take: "n < length xs \ last (xs \ Suc n) = xs ! n" by (simp add: last_nth length_greater_imp_not_empty min_eqR) corollary last_take2:" \ 0 < n; n \ length xs \ \ last (xs \ n) = xs ! (n - Suc 0)" apply (frule diff_Suc_less[THEN order_less_le_trans, of _ "length xs" 0], assumption) apply (drule last_take[of "n - Suc 0" xs]) apply simp done corollary nth_0_drop: "n \ length xs \ (xs \ n) ! 0 = xs ! n" by (cut_tac nth_drop[of n xs 0], simp+) lemma drop_eq_tl: "xs \ (Suc 0) = tl xs" by (simp add: drop_Suc) lemma drop_take_1: " n < length xs \ xs \ n \ (Suc 0) = [xs ! n]" by (simp add: take_hd hd_drop_conv_nth) lemma upt_append: "m \ n \ [0.. (xs @ ys) ! n = xs ! n" by (simp add: nth_append) lemma nth_append2: "length xs \ n \ (xs @ ys) ! n = ys ! (n - length xs)" by (simp add: nth_append) lemma list_all_conv: "list_all P xs = (\iys. (xs = ys) = (length xs = length ys \ (\i xs \ n = ys \ n; xs \ n = ys \ n \ \ xs = ys" apply (rule subst[OF append_take_drop_id[of n xs]]) apply (rule subst[OF append_take_drop_id[of n ys]]) apply simp done lemma list_take_drop_eq_conv: " (xs = ys) = (\n. (xs \ n = ys \ n \ xs \ n = ys \ n))" by (blast intro: list_take_drop_imp_eq) lemma list_take_eq_conv: "(xs = ys) = (\n. xs \ n = ys \ n)" apply (rule iffI, simp) apply (drule_tac x="max (length xs) (length ys)" in spec) apply simp done lemma list_drop_eq_conv: "(xs = ys) = (\n. xs \ n = ys \ n)" apply (rule iffI, simp) apply (drule_tac x=0 in spec) apply simp done abbreviation replicate' :: "'a \ nat \ 'a list" ("_\<^bsup>_\<^esup>" [1000,65]) where "x\<^bsup>n\<^esup> \ replicate n x" lemma replicate_snoc: "x\<^bsup>n\<^esup> @ [x] = x\<^bsup>Suc n\<^esup>" by (simp add: replicate_app_Cons_same) lemma eq_replicate_conv: "(\ilength xs\<^esup>)" apply (rule iffI) apply (simp add: expand_list_eq) apply clarsimp apply (rule ssubst[of xs "replicate (length xs) m"], assumption) apply (rule nth_replicate, simp) done lemma replicate_Cons_length: "length (x # a\<^bsup>n\<^esup>) = Suc n" by simp lemma replicate_pred_Cons_length: "0 < n \ length (x # a\<^bsup>n - Suc 0\<^esup>) = n" by simp lemma replicate_le_diff: "m \ n \ x\<^bsup>m\<^esup> @ x\<^bsup>n - m\<^esup> = x\<^bsup>n\<^esup>" by (simp add: replicate_add[symmetric]) lemma replicate_le_diff2: "\ k \ m; m \ n \ \ x\<^bsup>m - k\<^esup> @ x\<^bsup>n - m\<^esup> = x\<^bsup>n - k\<^esup>" by (subst replicate_add[symmetric], simp) lemma append_constant_length_induct_aux: "\xs. \ length xs div k = n; \ys. k = 0 \ length ys < k \ P ys; \xs ys. \ length xs = k; P ys \ \ P (xs @ ys) \ \ P xs" apply (case_tac "k = 0", blast) apply simp apply (induct n) apply (simp add: div_eq_0_conv') apply (subgoal_tac "k \ length xs") prefer 2 apply (rule div_gr_imp_gr_divisor[of 0], simp) apply (simp only: atomize_all atomize_imp, clarsimp) apply (erule_tac x="drop k xs" in allE) apply (simp add: div_diff_self2) apply (erule_tac x=undefined in allE) apply (erule_tac x="take k xs" in allE) apply (simp add: min_eqR) apply (erule_tac x="drop k xs" in allE) apply simp done lemma append_constant_length_induct: " \ \ys. k = 0 \ length ys < k \ P ys; \xs ys. \ length xs = k; P ys \ \ P (xs @ ys) \ \ P xs" by (simp add: append_constant_length_induct_aux[of _ _ "length xs div k"]) lemma zip_swap: "map (\(y,x). (x,y)) (zip ys xs) = (zip xs ys)" by (simp add: expand_list_eq) lemma zip_takeL: "(zip xs ys) \ n = zip (xs \ n) ys" by (simp add: expand_list_eq) lemma zip_takeR: "(zip xs ys) \ n = zip xs (ys \ n)" apply (subst zip_swap[of ys, symmetric]) apply (subst take_map) apply (subst zip_takeL) apply (simp add: zip_swap) done lemma zip_take: "(zip xs ys) \ n = zip (xs \ n) (ys \ n)" by (rule take_zip) lemma hd_zip: "\ xs \ []; ys \ [] \ \ hd (zip xs ys) = (hd xs, hd ys)" by (simp add: hd_conv_nth zip_not_empty_conv) lemma map_id: "map id xs = xs" by (simp add: id_def) lemma map_id_subst: "P (map id xs) \ P xs" by (subst map_id[symmetric]) lemma map_one: "map f [x] = [f x]" by simp lemma map_last: "xs \ [] \ last (map f xs) = f (last xs)" by (rule last_map) lemma filter_list_all: "list_all P xs \ filter P xs = xs" by (induct xs, simp+) lemma filter_snoc: "filter P (xs @ [x]) = (if P x then (filter P xs) @ [x] else filter P xs)" by (case_tac "P x", simp+) lemma filter_filter_eq: "list_all (\x. P x = Q x) xs \ filter P xs = filter Q xs" by (induct xs, simp+) lemma filter_nth: "\n. n < length (filter P xs) \ (filter P xs) ! n = xs ! (LEAST k. k < length xs \ n < card {i. i \ k \ i < length xs \ P (xs ! i)})" apply (induct xs rule: rev_induct, simp) apply (rename_tac x xs n) apply (simp only: filter_snoc) apply (simp split del: if_split) apply (case_tac "xs = []") apply (simp split del: if_split) apply (rule_tac t = "\k i. i = 0 \ i \ k \ P ([x] ! i)" and s = "\k i. i = 0 \ P x" in subst) apply (simp add: fun_eq_iff) apply fastforce apply (fastforce simp: Least_def) apply (rule_tac t = "\k. card {i. i \ k \ i < Suc (length xs) \ P ((xs @ [x]) ! i)}" and s = "\k. (card {i. i \ k \ i < length xs \ P (xs ! i)} + (if k \ length xs \ P x then Suc 0 else 0))" in subst) apply (clarsimp simp: fun_eq_iff split del: if_split, rename_tac k) apply (simp split del: if_split add: less_Suc_eq conj_disj_distribL conj_disj_distribR Collect_disj_eq) apply (subst card_Un_disjoint) apply (rule_tac n="length xs" in bounded_nat_set_is_finite, blast) apply (rule_tac n="Suc (length xs)" in bounded_nat_set_is_finite, blast) apply blast apply (rule_tac t = "\i. i < length xs \ P ((xs @ [x]) ! i)" and s = "\i. i < length xs \ P (xs ! i)" in subst) apply (rule fun_eq_iff[THEN iffD2]) apply (fastforce simp: nth_append1) - apply (rule nat_add_left_cancel[THEN iffD2]) + apply (rule add_left_cancel[THEN iffD2]) apply (rule_tac t = "\i. i = length xs \ i \ k \ P ((xs @ [x]) ! i)" and s = "\i. i = length xs \ i \ k \ P x" in subst) apply (rule fun_eq_iff[THEN iffD2]) apply fastforce apply (case_tac "length xs \ k") apply clarsimp apply (rule_tac t = "\i. i = length xs \ i \ k" and s = "\i. i = length xs" in subst) apply (rule fun_eq_iff[THEN iffD2]) apply fastforce apply simp apply simp apply (simp split del: if_split add: less_Suc_eq conj_disj_distribL conj_disj_distribR) apply (rule_tac t = "\k. k < length xs \ n < card {i. i \ k \ i < length xs \ P (xs ! i)} + (if length xs \ k \ P x then Suc 0 else 0)" and s = "\k. k < length xs \ n < card {i. i \ k \ i < length xs \ P (xs ! i)}" in subst) apply (simp add: fun_eq_iff) apply (rule_tac t = "\k. k = length xs \ n < card {i. i \ k \ i < length xs \ P (xs ! i)} + (if length xs \ k \ P x then Suc 0 else 0)" and s = "\k. k = length xs \ n < card {i. i \ k \ i < length xs \ P (xs ! i)} + (if P x then Suc 0 else 0)" in subst) apply (simp add: fun_eq_iff) apply (case_tac "n < length (filter P xs)") apply (rule_tac t = "(if P x then filter P xs @ [x] else filter P xs) ! n" and s = "(filter P xs) ! n" in subst) apply (simp add: nth_append1) apply (simp split del: if_split) apply (subgoal_tac "\k k \ i < length xs \ P (xs ! i)}") prefer 2 apply (rule_tac x="length xs - Suc 0" in exI) apply (simp add: length_filter_conv_card less_eq_le_pred[symmetric]) apply (subgoal_tac "\k\length xs. n < card {i. i \ k \ i < length xs \ P (xs ! i)}") prefer 2 apply (blast intro: less_imp_le) apply (subst Least_le_imp_le_disj) apply simp apply simp apply (rule sym, rule nth_append1) apply (rule LeastI2_ex, assumption) apply blast apply (simp add: linorder_not_less) apply (subgoal_tac "P x") prefer 2 apply (rule ccontr, simp) apply (simp add: length_snoc) apply (drule less_Suc_eq_le[THEN iffD1], drule_tac x=n in order_antisym, assumption) apply (simp add: nth_append2) apply (simp add: length_filter_conv_card) apply (rule_tac t = "\k. card {i. i < length xs \ P (xs ! i)} < card {i. i \ k \ i < length xs \ P (xs ! i)}" and s = "\k. False" in subst) apply (rule fun_eq_iff[THEN iffD2], rule allI, rename_tac k) apply (simp add: linorder_not_less) apply (rule card_mono) apply fastforce apply blast apply simp apply (rule_tac t = "(LEAST k. k = length xs \ card {i. i < length xs \ P (xs ! i)} < Suc (card {i. i \ k \ i < length xs \ P (xs ! i)}))" and s = "length xs" in subst) apply (rule sym, rule Least_equality) apply simp apply (rule le_imp_less_Suc) apply (rule card_mono) apply fastforce apply fastforce apply simp apply simp done subsubsection \Ordered lists\ fun list_ord :: "('a \ 'a \ bool) \ ('a::ord) list \ bool" where "list_ord ord (x1 # x2 # xs) = (ord x1 x2 \ list_ord ord (x2 # xs))" | "list_ord ord xs = True" definition list_asc :: "('a::ord) list \ bool" where "list_asc xs \ list_ord (\) xs" definition list_strict_asc :: "('a::ord) list \ bool" where "list_strict_asc xs \ list_ord (<) xs" value "list_asc [1::nat, 2, 2]" value "list_strict_asc [1::nat, 2, 2]" definition list_desc :: "('a::ord) list \ bool" where "list_desc xs \ list_ord (\) xs" definition list_strict_desc :: "('a::ord) list \ bool" where "list_strict_desc xs \ list_ord (>) xs" lemma list_ord_Nil: "list_ord ord []" by simp lemma list_ord_one: "list_ord ord [x]" by simp lemma list_ord_Cons: " list_ord ord (x # xs) = (xs = [] \ (ord x (hd xs) \ list_ord ord xs))" by (induct xs, simp+) lemma list_ord_Cons_imp: "\ list_ord ord xs; ord x (hd xs) \ \ list_ord ord (x # xs)" by (induct xs, simp+) lemma list_ord_append: "\ys. list_ord ord (xs @ ys) = (list_ord ord xs \ (ys = [] \ (list_ord ord ys \ (xs = [] \ ord (last xs) (hd ys)))))" apply (induct xs, fastforce) apply (case_tac xs, case_tac ys, fastforce+) done lemma list_ord_snoc: " list_ord ord (xs @ [x]) = (xs = [] \ (ord (last xs) x \ list_ord ord xs))" by (fastforce simp: list_ord_append) lemma list_ord_all_conv: " (list_ord ord xs) = (\n < length xs - 1. ord (xs ! n) (xs ! Suc n))" apply (rule iffI) apply (induct xs, simp) apply clarsimp apply (simp add: list_ord_Cons) apply (erule disjE, simp) apply clarsimp apply (case_tac n) apply (simp add: hd_conv_nth) apply simp apply (induct xs, simp) apply (simp add: list_ord_Cons) apply (case_tac "xs = []", simp) apply (drule meta_mp) apply (intro allI impI, rename_tac n) apply (drule_tac x="Suc n" in spec, simp) apply (drule_tac x=0 in spec) apply (simp add: hd_conv_nth) done lemma list_ord_imp: " \ \x y. ord x y \ ord' x y; list_ord ord xs \ \ list_ord ord' xs" apply (induct xs, simp) apply (simp add: list_ord_Cons) apply fastforce done corollary list_strict_asc_imp_list_asc: " list_strict_asc (xs::'a::preorder list) \ list_asc xs" by (unfold list_strict_asc_def list_asc_def, rule list_ord_imp[of "(<)"], rule order_less_imp_le) corollary list_strict_desc_imp_list_desc: " list_strict_desc (xs::'a::preorder list) \ list_desc xs" by (unfold list_strict_desc_def list_desc_def, rule list_ord_imp[of "(>)"], rule order_less_imp_le) lemma list_ord_trans_imp: "\i. \ transP ord; list_ord ord xs; j < length xs; i < j \ \ ord (xs ! i) (xs ! j)" apply (simp add: list_ord_all_conv) apply (induct j, simp) apply (case_tac "j < i", simp) apply (simp add: linorder_not_less) apply (case_tac "i = j", simp) apply (drule_tac x=i in meta_spec, simp) apply (drule_tac x=j in spec, simp add: Suc_less_pred_conv) apply (unfold trans_def) apply (drule_tac x="xs ! i" in spec, drule_tac x="xs ! j" in spec, drule_tac x="xs ! Suc j" in spec) apply simp done lemma list_ord_trans: " transP ord \ (list_ord ord xs) = (\j < length xs. \i < j. ord (xs ! i) (xs ! j))" apply (rule iffI) apply (simp add: list_ord_trans_imp) apply (simp add: list_ord_all_conv) done lemma list_ord_trans_refl_le: " \ transP ord; reflP ord \ \ (list_ord ord xs) = (\j < length xs. \i \ j. ord (xs ! i) (xs ! j))" apply (subst list_ord_trans, simp) apply (rule iffI) apply clarsimp apply (case_tac "i = j") apply (simp add: refl_on_def) apply simp+ done lemma list_ord_trans_refl_le_imp: " \ transP ord; \x y. ord x y \ ord' x y; reflP ord'; list_ord ord xs \ \ (\j < length xs. \i \ j. ord' (xs ! i) (xs ! j))" apply clarify apply (case_tac "i = j") apply (simp add: refl_on_def) apply (simp add: list_ord_trans_imp) done corollary list_asc_trans: " (list_asc (xs::'a::preorder list)) = (\j < length xs. \i < j. xs ! i \ xs ! j)" and list_strict_asc_trans: " (list_strict_asc (xs::'a::preorder list)) = (\j < length xs. \i < j. xs ! i < xs ! j)" and list_desc_trans: " (list_desc (xs::'a::preorder list)) = (\j < length xs. \i < j. xs ! j \ xs ! i)" and list_strict_desc_trans: " (list_strict_desc (xs::'a::preorder list)) = (\j < length xs. \i < j. xs ! j < xs ! i)" apply (unfold list_asc_def list_strict_asc_def list_desc_def list_strict_desc_def) apply (rule list_ord_trans, unfold trans_def, blast intro: order_trans order_less_trans)+ done corollary list_asc_trans_le: " (list_asc (xs::'a::preorder list)) = (\j < length xs. \i \ j. xs ! i \ xs ! j)" and list_desc_trans_le: " (list_desc (xs::'a::preorder list)) = (\j < length xs. \i \ j. xs ! j \ xs ! i)" apply (unfold list_asc_def list_strict_asc_def list_desc_def list_strict_desc_def) apply (rule list_ord_trans_refl_le, unfold trans_def, blast intro: order_trans, simp add: refl_on_def)+ done corollary list_strict_asc_trans_le: " (list_strict_asc (xs::'a::preorder list)) \ (\j < length xs. \i \ j. xs ! i \ xs ! j)" apply (unfold list_strict_asc_def) apply (rule list_ord_trans_refl_le_imp[where ord="(\)"]) apply (unfold trans_def, blast intro: order_trans) apply assumption apply (unfold refl_on_def, clarsimp) apply (rule list_ord_imp[where ord="(<)"], simp_all add: less_imp_le) done lemma list_ord_le_sorted_eq: "list_asc xs = sorted xs" apply (rule sym) apply (simp add: list_asc_def) apply (induct xs, simp) apply (rename_tac x xs) apply (simp add: list_ord_Cons) apply (case_tac "xs = []", simp_all) apply (case_tac "list_ord (\) xs", simp_all) apply (rule iffI) apply (drule_tac x="hd xs" in bspec, simp_all) apply clarify apply (drule in_set_conv_nth[THEN iffD1], clarsimp, rename_tac i1) apply (simp add: hd_conv_nth) apply (case_tac i1, simp) apply (rename_tac i2) apply simp apply (fold list_asc_def) apply (fastforce simp: list_asc_trans) done corollary list_asc_upto: "list_asc [m..n]" by (simp add: list_ord_le_sorted_eq) lemma list_strict_asc_upt: "list_strict_asc [m.. irrefl {(a, b). ord a b}; transP ord; list_ord ord xs; i < length xs; j < length xs; i < j \ \ xs ! i \ xs ! j" apply (subgoal_tac "\x y. ord x y \ x \ y") prefer 2 apply (rule ccontr) apply (simp add: irrefl_def) apply (simp add: list_ord_trans) done lemma list_ord_distinct: " \ irrefl {(a,b). ord a b}; transP ord; list_ord ord xs \ \ distinct xs" apply (simp add: distinct_conv_nth, intro allI impI, rename_tac i j) apply (drule neq_iff[THEN iffD1], erule disjE) apply (simp add: list_ord_distinct_aux) apply (simp add: list_ord_distinct_aux[THEN not_sym]) done lemma list_strict_asc_distinct: "list_strict_asc (xs::'a::preorder list) \ distinct xs" apply (rule_tac ord="(<)" in list_ord_distinct) apply (unfold irrefl_def list_strict_asc_def trans_def) apply (blast intro: less_trans)+ done lemma list_strict_desc_distinct: "list_strict_desc (xs::'a::preorder list) \ distinct xs" apply (rule_tac ord="(>)" in list_ord_distinct) apply (unfold irrefl_def list_strict_desc_def trans_def) apply (blast intro: less_trans)+ done subsubsection \Additional definitions and results for sublists\ primrec sublist_list :: "'a list \ nat list \ 'a list" where "sublist_list xs [] = []" | "sublist_list xs (y # ys) = (xs ! y) # (sublist_list xs ys)" value "sublist_list [0::int,10::int,20,30,40,50] [1::nat,2,3]" value "sublist_list [0::int,10::int,20,30,40,50] [1::nat,1,2,3]" value [nbe] "sublist_list [0::int,10::int,20,30,40,50] [1::nat,1,2,3,10]" lemma sublist_list_length: "length (sublist_list xs ys) = length ys" by (induct ys, simp_all) lemma sublist_list_append: " \zs. sublist_list xs (ys @ zs) = sublist_list xs ys @ sublist_list xs zs" by (induct ys, simp_all) lemma sublist_list_Nil: "sublist_list xs [] =[]" by simp lemma sublist_list_is_Nil_conv: " (sublist_list xs ys = []) = (ys = [])" apply (rule iffI) apply (rule ccontr) apply (clarsimp simp: neq_Nil_conv) apply simp done lemma sublist_list_eq_imp_length_eq: " sublist_list xs ys = sublist_list xs zs \ length ys = length zs" by (drule arg_cong[where f=length], simp add: sublist_list_length) lemma sublist_list_nth: " \n. n < length ys \ sublist_list xs ys ! n = xs ! (ys ! n)" apply (induct ys, simp) apply (case_tac n, simp_all) done lemma take_drop_eq_sublist_list: " m + n \ length xs \ xs \ m \ n = sublist_list xs [m.. nat list \ 'a list" where "sublist_list_if xs [] = []" | "sublist_list_if xs (y # ys) = (if y < length xs then (xs ! y) # (sublist_list_if xs ys) else (sublist_list_if xs ys))" value "sublist_list_if [0::int,10::int,20,30,40,50] [1::nat,2,3]" value "sublist_list_if [0::int,10::int,20,30,40,50] [1::nat,1,2,3]" value "sublist_list_if [0::int,10::int,20,30,40,50] [1::nat,1,2,3,10]" lemma sublist_list_if_sublist_list_filter_conv: "\xs. sublist_list_if xs ys = sublist_list xs (filter (\i. i < length xs) ys)" by (induct ys, simp+) corollary sublist_list_if_sublist_list_eq: "\xs. list_all (\i. i < length xs) ys \ sublist_list_if xs ys = sublist_list xs ys" by (simp add: sublist_list_if_sublist_list_filter_conv filter_list_all) corollary sublist_list_if_sublist_list_eq2: "\xs. \n sublist_list_if xs ys = sublist_list xs ys" by (rule sublist_list_if_sublist_list_eq, rule list_all_conv[THEN iffD2]) lemma sublist_list_if_Nil_left: "sublist_list_if [] ys = []" by (induct ys, simp+) lemma sublist_list_if_Nil_right: "sublist_list_if xs [] = []" by simp lemma sublist_list_if_length: " length (sublist_list_if xs ys) = length (filter (\i. i < length xs) ys)" by (simp add: sublist_list_if_sublist_list_filter_conv sublist_list_length) lemma sublist_list_if_append: " sublist_list_if xs (ys @ zs) = sublist_list_if xs ys @ sublist_list_if xs zs" by (simp add: sublist_list_if_sublist_list_filter_conv sublist_list_append) lemma sublist_list_if_snoc: " sublist_list_if xs (ys @ [y]) = sublist_list_if xs ys @ (if y < length xs then [xs ! y] else [])" by (simp add: sublist_list_if_append) lemma sublist_list_if_is_Nil_conv: " (sublist_list_if xs ys = []) = (list_all (\i. length xs \ i) ys)" by (simp add: sublist_list_if_sublist_list_filter_conv sublist_list_is_Nil_conv filter_empty_conv list_all_iff linorder_not_less) lemma sublist_list_if_nth: " n < length ((filter (\i. i < length xs) ys)) \ sublist_list_if xs ys ! n = xs ! ((filter (\i. i < length xs) ys) ! n)" by (simp add: sublist_list_if_sublist_list_filter_conv sublist_list_nth) lemma take_drop_eq_sublist_list_if: " m + n \ length xs \ xs \ m \ n = sublist_list_if xs [m..i\I. length xs \ i)" using [[hypsubst_thin = true]] by (fastforce simp: set_empty[symmetric] set_nths linorder_not_le[symmetric]) lemma nths_singleton2: "nths xs {y} = (if y < length xs then [xs ! y] else [])" apply (unfold nths_def) apply (induct xs rule: rev_induct, simp) apply (simp add: nth_append) done lemma nths_take_eq: " \ finite I; Max I < n \ \ nths (xs \ n) I = nths xs I" apply (case_tac "I = {}", simp) apply (case_tac "n < length xs") prefer 2 apply simp apply (rule_tac t = "nths xs I" and s = "nths (xs \ n @ xs \ n) I" in subst) apply simp apply (subst nths_append) apply (simp add: min_eqR) apply (rule_tac t="{j. j + n \ I}" and s="{}" in subst) apply blast apply simp done lemma nths_drop_eq: " n \ iMin I \ nths (xs \ n) {j. j + n \ I} = nths xs I" apply (case_tac "I = {}", simp) apply (case_tac "n < length xs") prefer 2 apply (simp add: nths_def filter_empty_conv linorder_not_less) apply (clarsimp, rename_tac a b) apply (drule set_zip_rightD) apply fastforce apply (rule_tac t = "nths xs I" and s = "nths (xs \ n @ xs \ n) I" in subst) apply simp apply (subst nths_append) apply (fastforce simp: nths_empty_conv min_eqR) done lemma nths_cut_less_eq: " length xs \ n \ nths xs (I \< n) = nths xs I" apply (simp add: nths_def cut_less_mem_iff) apply (rule_tac f="\xs. map fst xs" in arg_cong) apply (rule filter_filter_eq) apply (simp add: list_all_conv) done lemma nths_disjoint_Un: " \ finite A; Max A < iMin B \ \ nths xs (A \ B) = nths xs A @ nths xs B" apply (case_tac "A = {}", simp) apply (case_tac "B = {}", simp) apply (case_tac "length xs \ iMin B") apply (subst nths_cut_less_eq[of xs "iMin B", symmetric], assumption) apply (simp (no_asm_simp) add: cut_less_Un cut_less_Min_empty cut_less_Max_all) apply (simp add: nths_empty_conv iMin_ge_iff) apply (simp add: linorder_not_le) apply (rule_tac t = "nths xs (A \ B)" and s = "nths (xs \ (iMin B) @ xs \ (iMin B)) (A \ B)" in subst) apply simp apply (subst nths_append) apply (simp add: min_eqR) apply (subst nths_cut_less_eq[where xs="xs \ iMin B" and n="iMin B", symmetric], simp) apply (simp add: cut_less_Un cut_less_Min_empty cut_less_Max_all) apply (simp add: nths_take_eq) apply (rule_tac t = "\j. j + iMin B \ A \ j + iMin B \ B" and s = "\j. j + iMin B \ B" in subst) apply (force simp: fun_eq_iff) apply (simp add: nths_drop_eq) done corollary nths_disjoint_insert_left: " \ finite I; x < iMin I \ \ nths xs (insert x I) = nths xs {x} @ nths xs I" apply (rule_tac t="insert x I" and s="{x} \ I" in subst, simp) apply (subst nths_disjoint_Un) apply simp_all done corollary nths_disjoint_insert_right: " \ finite I; Max I < x \ \ nths xs (insert x I) = nths xs I @ nths xs {x}" apply (rule_tac t="insert x I" and s="I \ {x}" in subst, simp) apply (subst nths_disjoint_Un) apply simp_all done lemma nths_all: "{.. I \ nths xs I = xs" apply (case_tac "xs = []", simp) apply (rule_tac t = "I" and s = "I \< (length xs) \ I \\ (length xs)" in subst) apply (simp add: cut_less_cut_ge_ident) apply (rule_tac t = "I \< length xs" and s = "{..\ (length xs) = {}", simp) apply (subst nths_disjoint_Un[OF finite_lessThan]) apply (rule less_imp_Max_less_iMin[OF finite_lessThan]) apply blast apply blast apply (blast intro: less_le_trans) apply (fastforce simp: nths_empty_conv) done corollary nths_UNIV: "nths xs UNIV = xs" by (rule nths_all[OF subset_UNIV]) lemma sublist_list_nths_eq: "\xs. list_strict_asc ys \ sublist_list_if xs ys = nths xs (set ys)" apply (case_tac "xs = []") apply (simp add: sublist_list_if_Nil_left) apply (induct ys rule: rev_induct, simp) apply (rename_tac y ys xs) apply (case_tac "ys = []") apply (simp add: nths_singleton2) apply (unfold list_strict_asc_def) apply (simp add: sublist_list_if_snoc split del: if_split) apply (frule list_ord_append[THEN iffD1]) apply (clarsimp split del: if_split) apply (subst nths_disjoint_insert_right) apply simp apply (clarsimp simp: in_set_conv_nth, rename_tac i) apply (drule_tac i=i and j="length ys" in list_strict_asc_trans[unfolded list_strict_asc_def, THEN iffD1, rule_format]) apply (simp add: nth_append split del: if_split)+ apply (simp add: nths_singleton2) done lemma set_sublist_list_if: "\xs. set (sublist_list_if xs ys) = {xs ! i |i. i < length xs \ i \ set ys}" apply (induct ys, simp_all) apply blast done lemma set_sublist_list: " list_all (\i. i < length xs) ys \ set (sublist_list xs ys) = {xs ! i |i. i < length xs \ i \ set ys}" by (simp add: sublist_list_if_sublist_list_eq[symmetric] set_sublist_list_if) lemma set_sublist_list_if_eq_set_sublist: "set (sublist_list_if xs ys) = set (nths xs (set ys))" by (simp add: set_nths set_sublist_list_if) lemma set_sublist_list_eq_set_sublist: " list_all (\i. i < length xs) ys \ set (sublist_list xs ys) = set (nths xs (set ys))" by (simp add: sublist_list_if_sublist_list_eq[symmetric] set_sublist_list_if_eq_set_sublist) subsubsection \Natural set images with lists\ definition f_image :: "'a list \ nat set \ 'a set" (infixr "`\<^sup>f" 90) where "xs `\<^sup>f A \ {y. \n\A. n < length xs \ y = xs ! n}" abbreviation f_range :: "'a list \ 'a set" where "f_range xs \ f_image xs UNIV" lemma f_image_eqI[simp, intro]: " \ x = xs ! n; n \ A; n < length xs \ \ x \ xs `\<^sup>f A" by (unfold f_image_def, blast) lemma f_imageI: "\ n \ A; n < length xs \ \ xs ! n \ xs `\<^sup>f A" by blast lemma rev_f_imageI: "\ n \ A; n < length xs; x = xs ! n \ \ x \ xs `\<^sup>f A" by (rule f_image_eqI) lemma f_imageE[elim!]: " \ x \ xs `\<^sup>f A; \n. \ x = xs ! n; n \ A; n < length xs \ \ P \ \ P" by (unfold f_image_def, blast) lemma f_image_Un: "xs `\<^sup>f (A \ B) = xs `\<^sup>f A \ xs `\<^sup>f B" by blast lemma f_image_mono: "A \ B ==> xs `\<^sup>f A \ xs `\<^sup>f B" by blast lemma f_image_iff: "(x \ xs `\<^sup>f A) = (\n\A. n < length xs \ x = xs ! n)" by blast lemma f_image_subset_iff: " (xs `\<^sup>f A \ B) = (\n\A. n < length xs \ xs ! n \ B)" by blast lemma subset_f_image_iff: "(B \ xs `\<^sup>f A) = (\A'\A. B = xs `\<^sup>f A')" apply (rule iffI) apply (rule_tac x="{ n. n \ A \ n < length xs \ xs ! n \ B }" in exI) apply blast apply (blast intro: f_image_mono) done lemma f_image_subsetI: " \ \n. n \ A \ n < length xs \ xs ! n \ B \ \ xs `\<^sup>f A \ B" by blast lemma f_image_empty: "xs `\<^sup>f {} = {}" by blast lemma f_image_insert_if: " xs `\<^sup>f (insert n A) = ( if n < length xs then insert (xs ! n) (xs `\<^sup>f A) else (xs `\<^sup>f A))" by (split if_split, blast) lemma f_image_insert_eq1: " n < length xs \ xs `\<^sup>f (insert n A) = insert (xs ! n) (xs `\<^sup>f A)" by (simp add: f_image_insert_if) lemma f_image_insert_eq2: " length xs \ n \ xs `\<^sup>f (insert n A) = (xs `\<^sup>f A)" by (simp add: f_image_insert_if) lemma insert_f_image: " \ n \ A; n < length xs \ \ insert (xs ! n) (xs `\<^sup>f A) = (xs `\<^sup>f A)" by blast lemma f_image_is_empty: "(xs `\<^sup>f A = {}) = ({x. x \ A \ x < length xs} = {})" by blast lemma f_image_Collect: "xs `\<^sup>f {n. P n} = {xs ! n |n. P n \ n < length xs}" by blast lemma f_image_eq_set: "\n A \ xs `\<^sup>f A = set xs" by (fastforce simp: in_set_conv_nth) lemma f_range_eq_set: "f_range xs = set xs" by (simp add: f_image_eq_set) lemma f_image_eq_set_nths: "xs `\<^sup>f A = set (nths xs A)" by (unfold set_nths, blast) lemma f_image_eq_set_sublist_list_if: "xs `\<^sup>f (set ys) = set (sublist_list_if xs ys)" by (simp add: set_sublist_list_if_eq_set_sublist f_image_eq_set_nths) lemma f_image_eq_set_sublist_list: " list_all (\i. i < length xs) ys \ xs `\<^sup>f (set ys) = set (sublist_list xs ys)" by (simp add: sublist_list_if_sublist_list_eq f_image_eq_set_sublist_list_if) lemma f_range_eqI: "\ x = xs ! n; n < length xs \ \ x \ f_range xs" by blast lemma f_rangeI: "n < length xs \ xs ! n \ f_range xs" by blast lemma f_rangeE[elim?]: " \ x \ f_range xs; \n. \ n < length xs; x = xs ! n \ \ P \ \ P" by blast subsubsection \Mapping lists of functions to lists\ primrec map_list :: "('a \ 'b) list \ 'a list \ 'b list" where "map_list [] xs = []" | "map_list (f # fs) xs = f (hd xs) # map_list fs (tl xs)" lemma map_list_Nil: "map_list [] xs = []" by simp lemma map_list_Cons_Cons: " map_list (f # fs) (x # xs) = (f x) # map_list fs xs" by simp lemma map_list_length: "\xs. length (map_list fs xs) = length fs" by (induct fs, simp+) corollary map_list_empty_conv: " (map_list fs xs = []) = (fs = [])" by (simp del: length_0_conv add: length_0_conv[symmetric] map_list_length) corollary map_list_not_empty_conv: " (map_list fs xs \ []) = (fs \ [])" by (simp add: map_list_empty_conv) lemma map_list_nth: "\n xs. \ n < length fs; n < length xs \ \ (map_list fs xs ! n) = (fs ! n) (xs ! n)" apply (induct fs, simp+) apply (case_tac n) apply (simp add: hd_conv_nth) apply (simp add: nth_tl_eq_nth_Suc2) done lemma map_list_xs_take: "\n xs. length fs \ n \ map_list fs (xs \ n) = map_list fs xs" apply (induct fs, simp+) apply (rename_tac fs n xs) apply (simp add: tl_take) done lemma map_list_take: "\n xs. (map_list fs xs) \ n = (map_list (fs \ n) xs)" apply (induct fs, simp) apply (case_tac n, simp+) done lemma map_list_take_take: "\n xs. (map_list fs xs) \ n = (map_list (fs \ n) (xs \ n))" by (simp add: map_list_take map_list_xs_take) lemma map_list_drop: "\n xs. (map_list fs xs) \ n = (map_list (fs \ n) (xs \ n))" apply (induct fs, simp) apply (case_tac n) apply (simp add: drop_Suc)+ done lemma map_list_append_append: "\xs1 . length fs1 = length xs1 \ map_list (fs1 @ fs2) (xs1 @ xs2) = map_list fs1 xs1 @ map_list fs2 xs2" apply (induct fs1, simp+) apply (case_tac "xs1", simp+) done lemma map_list_snoc_snoc: " length fs = length xs \ map_list (fs @ [f]) (xs @ [x]) = map_list fs xs @ [f x]" by (simp add: map_list_append_append) lemma map_list_snoc: "\xs. length fs < length xs \ map_list (fs @ [f]) xs = map_list fs xs @ [f (xs ! (length fs))]" apply (induct fs) apply (simp add: hd_conv_nth) apply (simp add: nth_tl_eq_nth_Suc2) done lemma map_list_Cons_if: " map_list fs (x # xs) = (if (fs = []) then [] else ( ((hd fs) x) # map_list (tl fs) xs))" by (case_tac "fs", simp+) lemma map_list_Cons_not_empty: " fs \ [] \ map_list fs (x # xs) = ((hd fs) x) # map_list (tl fs) xs" by (simp add: map_list_Cons_if) lemma map_eq_map_list_take: "\xs. \ length fs \ length xs; list_all (\x. x = f) fs \ \ map_list fs xs = map f (xs \ length fs)" apply (induct fs, simp+) apply (case_tac xs, simp+) done lemma map_eq_map_list_take2: " \ length fs = length xs; list_all (\x. x = f) fs \ \ map_list fs xs = map f xs" by (simp add: map_eq_map_list_take) lemma map_eq_map_list_replicate: " map_list (f\<^bsup>length xs\<^esup>) xs = map f xs" by (induct xs, simp+) subsubsection \Mapping functions with two arguments to lists\ primrec map2 :: " \ \Function taking two parameters\ ('a \ 'b \ 'c) \ \ \Lists of parameters\ 'a list \ 'b list \ 'c list" where "map2 f [] ys = []" | "map2 f (x # xs) ys = f x (hd ys) # map2 f xs (tl ys)" lemma map2_map_list_conv: "\ys. map2 f xs ys = map_list (map f xs) ys" by (induct xs, simp+) lemma map2_Nil: "map2 f [] ys = []" by simp lemma map2_Cons_Cons: " map2 f (x # xs) (y # ys) = (f x y) # map2 f xs ys" by simp lemma map2_length: "\ys. length (map2 f xs ys) = length xs" by (induct xs, simp+) corollary map2_empty_conv: " (map2 f xs ys = []) = (xs = [])" by (simp del: length_0_conv add: length_0_conv[symmetric] map2_length) corollary map2_not_empty_conv: " (map2 f xs ys \ []) = (xs \ [])" by (simp add: map2_empty_conv) lemma map2_nth: "\n ys. \ n < length xs; n < length ys \ \ (map2 f xs ys ! n) = f (xs ! n) (ys ! n)" by (simp add: map2_map_list_conv map_list_nth) lemma map2_ys_take: "\n ys. length xs \ n \ map2 f xs (ys \ n) = map2 f xs ys" by (simp add: map2_map_list_conv map_list_xs_take) lemma map2_take: "\n ys. (map2 f xs ys) \ n = (map2 f (xs \ n) ys)" by (simp add: map2_map_list_conv take_map map_list_take) lemma map2_take_take: "\n ys. (map2 f xs ys) \ n = (map2 f (xs \ n) (ys \ n))" by (simp add: map2_take map2_ys_take) lemma map2_drop: "\n ys. (map2 f xs ys) \ n = (map2 f (xs \ n) (ys \ n))" by (simp add: map2_map_list_conv map_list_drop drop_map) lemma map2_append_append: "\ys1 . length xs1 = length ys1 \ map2 f (xs1 @ xs2) (ys1 @ ys2) = map2 f xs1 ys1 @ map2 f xs2 ys2" by (simp add: map2_map_list_conv map_list_append_append) lemma map2_snoc_snoc: " length xs = length ys \ map2 f (xs @ [x]) (ys @ [y]) = map2 f xs ys @ [f x y]" by (simp add: map2_append_append) lemma map2_snoc: "\ys. length xs < length ys \ map2 f (xs @ [x]) ys = map2 f xs ys @ [f x (ys ! (length xs))]" by (simp add: map2_map_list_conv map_list_snoc) lemma map2_Cons_if: " map2 f xs (y # ys) = (if (xs = []) then [] else ( (f (hd xs) y) # map2 f (tl xs) ys))" by (case_tac "xs", simp+) lemma map2_Cons_not_empty: " xs \ [] \ map2 f xs (y # ys) = (f (hd xs) y) # map2 f (tl xs) ys" by (simp add: map2_Cons_if) lemma map2_append1_take_drop: " length xs1 \ length ys \ map2 f (xs1 @ xs2) ys = map2 f xs1 (ys \ length xs1) @ map2 f xs2 (ys \ length xs1)" apply (rule_tac t = "map2 f (xs1 @ xs2) ys" and s = "map2 f (xs1 @ xs2) (ys \ length xs1 @ ys \ length xs1)" in subst) apply simp apply (simp add: map2_append_append del: append_take_drop_id) done lemma map2_append2_take_drop: " length ys1 \ length xs \ map2 f xs (ys1 @ ys2) = map2 f (xs \ length ys1) ys1 @ map2 f (xs \ length ys1) ys2" apply (rule_tac t = "map2 f xs (ys1 @ ys2)" and s = "map2 f (xs \ length ys1 @ xs \ length ys1) (ys1 @ ys2)" in subst) apply simp apply (simp add: map2_append_append del: append_take_drop_id) done lemma map2_cong: " \ xs1 = xs2; ys1 = ys2; length xs2 \ length ys2; \x y. \ x \ set xs2; y \ set ys2 \ \ f x y = g x y \ \ map2 f xs1 ys1 = map2 g xs2 ys2" by (simp (no_asm_simp) add: expand_list_eq map2_length map2_nth) lemma map2_eq_conv: " length xs \ length ys \ (map2 f xs ys = map2 g xs ys) = (\in\<^esup> y\<^bsup>n\<^esup> = (f x y)\<^bsup>n\<^esup>" by (induct n, simp+) lemma map2_zip_conv: "\ys. length xs \ length ys \ map2 f xs ys = map (\(x,y). f x y) (zip xs ys)" apply (induct xs, simp) apply (case_tac ys, simp+) done lemma map2_rev: "\ys. length xs = length ys \ rev (map2 f xs ys) = map2 f (rev xs) (rev ys)" apply (induct xs, simp) apply (case_tac ys, simp) apply (simp add: map2_Cons_Cons map2_snoc_snoc) done hide_const (open) map2 end diff --git a/thys/LocalLexing/Derivations.thy b/thys/LocalLexing/Derivations.thy --- a/thys/LocalLexing/Derivations.thy +++ b/thys/LocalLexing/Derivations.thy @@ -1,942 +1,942 @@ theory Derivations imports CFG ListTools InductRules begin (* leftderives and leftderives1 *) context CFG begin lemma [simp]: "is_terminal t \ is_symbol t" by (auto simp add: is_terminal_def is_symbol_def) lemma [simp]: "is_sentence []" by (auto simp add: is_sentence_def) lemma [simp]: "is_word []" by (auto simp add: is_word_def) lemma [simp]: "is_word u \ is_sentence u" by (induct u, auto simp add: is_word_def is_sentence_def) definition leftderives1 :: "sentence \ sentence \ bool" where "leftderives1 u v = (\ x y N \. u = x @ [N] @ y \ v = x @ \ @ y \ is_word x \ is_sentence y \ (N, \) \ \)" lemma leftderives1_implies_derives1[simp]: "leftderives1 u v \ derives1 u v" apply (auto simp add: leftderives1_def derives1_def) apply (rule_tac x="x" in exI) apply (rule_tac x="y" in exI) apply (rule_tac x="N" in exI) by auto definition leftderivations1 :: "(sentence \ sentence) set" where "leftderivations1 = { (u,v) | u v. leftderives1 u v }" lemma [simp]: "leftderivations1 \ derivations1" by (auto simp add: leftderivations1_def derivations1_def) definition leftderivations :: "(sentence \ sentence) set" where "leftderivations = leftderivations1^*" lemma rtrancl_subset_implies: "a \ b \ a \ b^*" by blast lemma leftderivations_subset_derivations[simp]: "leftderivations \ derivations" apply (simp add: leftderivations_def derivations_def) apply (rule rtrancl_subset_rtrancl) apply (rule rtrancl_subset_implies) by simp definition leftderives :: "sentence \ sentence \ bool" where "leftderives u v = ((u, v) \ leftderivations)" lemma leftderives_implies_derives[simp]: "leftderives u v \ derives u v" apply (auto simp add: leftderives_def derives_def) by (rule subsetD[OF leftderivations_subset_derivations]) definition is_leftderivation :: "sentence \ bool" where "is_leftderivation u = leftderives [\] u" lemma leftderivation_implies_derivation[simp]: "is_leftderivation u \ is_derivation u" by (simp add: is_leftderivation_def is_derivation_def) lemma leftderives_refl[simp]: "leftderives u u" by (auto simp add: leftderives_def leftderivations_def) lemma leftderives1_implies_leftderives[simp]:"leftderives1 a b \ leftderives a b" by (auto simp add: leftderives_def leftderivations_def leftderivations1_def) lemma leftderives_trans: "leftderives a b \ leftderives b c \ leftderives a c" by (auto simp add: leftderives_def leftderivations_def) lemma leftderives1_eq_leftderivations1: "leftderives1 x y = ((x, y) \ leftderivations1)" by (simp add: leftderivations1_def) lemma leftderives_induct[consumes 1, case_names Base Step]: assumes derives: "leftderives a b" assumes Pa: "P a" assumes induct: "\y z. leftderives a y \ leftderives1 y z \ P y \ P z" shows "P b" proof - note rtrancl_lemma = rtrancl_induct[where a = a and b = b and r = leftderivations1 and P=P] from derives Pa induct rtrancl_lemma show "P b" by (metis leftderives_def leftderivations_def leftderives1_eq_leftderivations1) qed end (* Basic lemmas about derives1 and derives *) context CFG begin 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 end (* Derives1 and Derivation, LDerives1 and LDerivation *) context CFG begin definition Derives1 :: "sentence \ nat \ rule \ sentence \ bool" where "Derives1 u i r v = (\ x y N \. u = x @ [N] @ y \ v = x @ \ @ y \ is_sentence x \ is_sentence y \ (N, \) \ \ \ 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) lemma Derives1_unique_dest: "Derives1 u i r v \ Derives1 u i r w \ v = w" by (auto simp add: Derives1_def derives1_def) lemma Derives1_unique_src: "Derives1 u i r w \ Derives1 v i r w \ u = v" by (auto simp add: Derives1_def derives1_def) type_synonym derivation = "(nat \ rule) list" fun Derivation :: "sentence \ derivation \ 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 auto 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 Derives1_take: "Derives1 a i r b \ take i a = take i b" by (auto simp add: Derives1_def) lemma Derives1_drop: "Derives1 a i r b \ drop (Suc i) a = drop (i + length (snd r)) b" by (auto simp add: Derives1_def) lemma Derives1_bound: "Derives1 a i r b \ i < length a" by (auto simp add: Derives1_def) lemma Derives1_length: "Derives1 a i r b \ length b = length a + length (snd r) - 1" by (auto simp add: Derives1_def) definition leftmost :: "nat \ sentence \ bool" where "leftmost i s = (i < length s \ is_word (take i s) \ is_nonterminal (s ! i))" lemma set_take: "set (take n s) = { s ! i | i. i < n \ i < length s}" proof (cases "n \ length s") case True thus ?thesis by (subst List.nth_image[symmetric], auto) next case False thus ?thesis apply (subst set_conv_nth) by (metis less_trans linear not_le take_all) qed lemma list_all_take: "list_all P (take n s) = (\ i. i < n \ i < length s \ P (s ! i))" by (auto simp add: set_take list_all_iff) lemma is_sentence_concat: "is_sentence (x@y) = (is_sentence x \ is_sentence y)" by (auto simp add: is_sentence_def) lemma is_sentence_cons: "is_sentence (x#xs) = (is_symbol x \ is_sentence xs)" by (auto simp add: is_sentence_def) lemma rule_nonterminal_type[simp]: "(N, \) \ \ \ is_nonterminal N" apply (insert validRules) by (auto simp add: is_nonterminal_def) lemma rule_\_type[simp]: "(N, \) \ \ \ is_sentence \" apply (insert validRules) by (auto simp add: is_sentence_def is_symbol_def list_all_iff is_terminal_def is_nonterminal_def) lemma [simp]: "is_nonterminal N \ is_symbol N" by (simp add: is_symbol_def) lemma Derives1_sentence1[elim]: "Derives1 a i r b \ is_sentence a" by (auto simp add: Derives1_def is_sentence_cons is_sentence_concat) lemma Derives1_sentence2[elim]: "Derives1 a i r b \ is_sentence b" by (auto simp add: Derives1_def is_sentence_cons is_sentence_concat) lemma [elim]: "Derives1 a i r b \ r \ \" using Derives1_def by auto lemma is_sentence_symbol: "is_sentence a \ i < length a \ is_symbol (a ! i)" by (simp add: is_sentence_def list_all_iff) lemma is_symbol_distinct: "is_symbol x \ is_terminal x \ is_nonterminal x" apply (insert disjunct_symbols) apply (auto simp add: is_symbol_def is_terminal_def is_nonterminal_def) done lemma is_terminal_nonterminal: "is_terminal x \ is_nonterminal x \ False" by (metis is_symbol_def is_symbol_distinct) lemma Derives1_leftmost: assumes "Derives1 a i r b" shows "\ j. leftmost j a \ j \ i" proof - let ?J = "{j . j < length a \ is_nonterminal (a ! j)}" let ?M = "Min ?J" from assms have J1:"i \ ?J" apply (auto simp add: Derives1_def is_nonterminal_def) by (metis (mono_tags, lifting) prod.simps(2) validRules) have J2:"finite ?J" by auto note J = J1 J2 from J have M1: "?M \ ?J" by (rule_tac Min_in, auto) { fix j assume "j \ ?J" with J have "?M \ j" by auto } note M3 = this[OF J1] from assms have a_sentence: "is_sentence a" by (simp add: Derives1_sentence1) have is_word: "is_word (take ?M a)" apply (auto simp add: is_word_def list_all_take) proof - fix i assume i_less_M: "i < ?M" assume i_inbounds: "i < length a" show "is_terminal (a ! i)" proof(cases "is_terminal (a ! i)") case True thus ?thesis by auto next case False then have "is_nonterminal (a ! i)" using i_inbounds a_sentence is_sentence_symbol is_symbol_distinct by blast then have "i \ ?J" by (metis i_inbounds mem_Collect_eq) then have "?M < i" by (metis J2 Min_le i_less_M leD) then have "False" by (metis i_less_M less_asym') then show ?thesis by auto qed qed show ?thesis apply (rule_tac exI[where x="?M"]) apply (simp add: leftmost_def) by (metis (mono_tags, lifting) M1 M3 is_word mem_Collect_eq) qed lemma Derivation_leftmost: "D \ [] \ Derivation a D b \ \ i. leftmost i a" apply (case_tac "D") apply (auto) apply (metis Derives1_leftmost) done lemma nonword_has_nonterminal: "is_sentence a \ \ (is_word a) \ \ k. k < length a \ is_nonterminal (a ! k)" apply (auto simp add: is_sentence_def list_all_iff is_word_def) by (metis in_set_conv_nth is_symbol_distinct) lemma leftmost_cons_nonterminal: "is_nonterminal x \ leftmost 0 (x#xs)" by (metis CFG.is_word_def CFG_axioms leftmost_def length_greater_0_conv list.distinct(1) list_all_simps(2) nth_Cons_0 take_Cons') lemma leftmost_cons_terminal: "is_terminal x \ leftmost i (x#xs) = (i > 0 \ leftmost (i - 1) xs)" by (metis Suc_diff_1 Suc_less_eq is_terminal_nonterminal is_word_def leftmost_def length_Cons list_all_simps(1) not_gr0 nth_Cons' take_Cons') lemma is_nonterminal_cons_terminal: "is_terminal x \ k < length (x # a) \ is_nonterminal ((x # a) ! k) \ k > 0 \ k - 1 < length a \ is_nonterminal (a ! (k - 1))" by (metis One_nat_def Suc_leI is_terminal_nonterminal less_diff_conv2 list.size(4) nth_non_equal_first_eq) lemma leftmost_exists: "is_sentence a \ k < length a \ is_nonterminal (a ! k) \ \ i. leftmost i a \ i \ k" proof (induct a arbitrary: k) case Nil thus ?case by auto next case (Cons x a) show ?case proof(cases "is_nonterminal x") case True thus ?thesis apply(rule_tac exI[where x="0"]) apply (simp add: leftmost_cons_nonterminal) done next case False then have x: "is_terminal x" by (metis Cons.prems(1) is_sentence_cons is_symbol_distinct) note k = is_nonterminal_cons_terminal[OF x Cons(3) Cons(4)] with Cons have "\i. leftmost i a \ i \ k - 1" by (metis is_sentence_cons) then show ?thesis apply (auto simp add: leftmost_cons_terminal[OF x]) - by (metis Nat.le_diff_conv2 Suc_leI add_Suc_right add_diff_cancel_right' k + by (metis le_diff_conv2 Suc_leI add_Suc_right add_diff_cancel_right' k le_0_eq le_imp_less_Suc nat_le_linear) qed qed lemma nonword_leftmost_exists: "is_sentence a \ \ (is_word a) \ \ i. leftmost i a" by (metis leftmost_exists nonword_has_nonterminal) lemma leftmost_unaffected_Derives1: "leftmost j a \ j < i \ Derives1 a i r b \ leftmost j b" apply (simp add: leftmost_def) proof - assume a1: "j < length a \ is_word (take j a) \ is_nonterminal (a ! j)" assume a2: "j < i" assume "Derives1 a i r b" then have f3: "take i a = take i b" by (metis Derives1_take) have f4: "\n ss ssa. take (length (take n (ss::symbol list))) (ssa::symbol list) = take (length ss) (take n ssa)" by (metis length_take take_take) have f5: "\ss. take j (ss::symbol list) = take i (take j ss)" using a2 by (metis dual_order.strict_implies_order min.absorb2 take_take) have f6: "length (take j a) = j" using a1 by (metis dual_order.strict_implies_order length_take min.absorb2) then have f7: "\n. min j n = length (take n (take j a))" by (metis length_take) have f8: "\n ss. n = length (take n (ss::symbol list)) \ length ss < n" by (metis leI length_take min.absorb2) have f9: "\ss. take j (ss::symbol list) = take j (take i ss)" using f7 f6 f5 by (metis take_take) have f10: "\ss n. length (ss::symbol list) \ n \ length (take n ss) = n" using f8 by (metis dual_order.strict_implies_order) then have f11: "\ss ssa. length (ss::symbol list) = length (take (length ss) (ssa::symbol list)) \ length (take (length ssa) ss) = length ssa" by (metis length_take min.absorb2) have f12: "\ss ssa n. take (length (ss::symbol list)) (ssa::symbol list) = take n (take (length ss) ssa) \ length (take n ss) = n" using f10 by (metis min.absorb2 take_take) { assume "\ j < j" { assume "\ j < j \ i \ j" moreover { assume "length a \ j \ length (take i a) \ i" then have "\ss. length (take (length (take i (take (length a) (ss::symbol list)))) (take j ss)) \ length (take i (take (length a) ss))" using f12 f11 f6 f5 f4 by metis then have "\ss ssa. take (length (ss::symbol list)) (take j (ssa::symbol list)) \ take (length ss) (take i (take (length a) ssa))" using f11 by metis then have "length b \ j" using f9 f4 f3 by metis } ultimately have "length b \ j" using f7 f6 f5 f3 a1 by (metis length_take) } then have "length b = j \ j < j" using a2 by metis } then have "j < length b" using f9 f8 f7 f6 f4 f3 by metis then show "j < length b \ is_word (take j b) \ is_nonterminal (b ! j)" using f9 f3 a2 a1 by (metis nth_take) qed definition derivation_ge :: "derivation \ nat \ bool" where "derivation_ge D i = (\ d \ set D. fst d \ i)" lemma derivation_ge_cons: "derivation_ge (d#D) i = (fst d \ i \ derivation_ge D i)" by (auto simp add: derivation_ge_def) lemma derivation_ge_append: "derivation_ge (D@E) i = (derivation_ge D i \ derivation_ge E i)" by (auto simp add: derivation_ge_def) lemma leftmost_unaffected_Derivation: "derivation_ge D (Suc i) \ leftmost i a \ Derivation a D b \ leftmost i b" proof (induct D arbitrary: a) case Nil thus ?case by auto next case (Cons d D) then have "\ x. Derives1 a (fst d) (snd d) x \ Derivation x D b" by simp then obtain x where x1: "Derives1 a (fst d) (snd d) x" and x2: "Derivation x D b" by blast with Cons have leftmost_x: "leftmost i x" apply (rule_tac leftmost_unaffected_Derives1[ where a=a and j=i and b="x" and i="fst d" and r="snd d"]) by (auto simp add: derivation_ge_def) with Cons x2 show ?case by (auto simp add: derivation_ge_def) qed lemma le_Derives1_take: assumes le: "i \ j" and D: "Derives1 a j r b" shows "take i a = take i b" proof - note Derives1_take[where a=a and i=j and r=r and b=b] with le D show ?thesis by (rule_tac le_take_same[where i=i and j=j], auto) qed lemma Derivation_take: "derivation_ge D i \ Derivation a D b \ take i a = take i b" proof(induct D arbitrary: a b) case Nil thus ?case by auto next case (Cons d D) then have "\ x. Derives1 a (fst d) (snd d) x \ Derivation x D b" by simp then obtain x where ax: "Derives1 a (fst d) (snd d) x" and xb: "Derivation x D b" by blast from derivation_ge_cons Cons(2) have d: "i \ fst d" and D: "derivation_ge D i" by auto note take_i_xb = Cons(1)[OF D xb] note take_i_ax = le_Derives1_take[OF d ax] from take_i_xb take_i_ax show ?case by auto qed lemma leftmost_cons_less: "i < length u \ leftmost i (u@v) = leftmost i u" by (auto simp add: leftmost_def nth_append) lemma leftmost_is_nonterminal: "leftmost i u \ is_nonterminal (u ! i)" by (metis leftmost_def) lemma is_word_is_terminal: "i < length u \ is_word u \ is_terminal (u ! i)" by (metis is_word_def list_all_length) lemma leftmost_append: assumes leftmost: "leftmost i (u@v)" and is_word: "is_word u" shows "length u \ i" proof (cases "i < length u") case False thus ?thesis by auto next case True with leftmost have "leftmost i u" using leftmost_cons_less[OF True] by simp then have is_nonterminal: "is_nonterminal (u ! i)" by (rule leftmost_is_nonterminal) note is_terminal = is_word_is_terminal[OF True is_word] note is_terminal_nonterminal[OF is_terminal is_nonterminal] then show ?thesis by auto qed lemma derivation_ge_empty[simp]: "derivation_ge [] i" by (simp add: derivation_ge_def) lemma leftmost_notword: "leftmost i a \ j > i \ \ (is_word (take j a))" by (metis is_terminal_nonterminal is_word_def leftmost_def list_all_take) lemma leftmost_unique: "leftmost i a \ leftmost j a \ i = j" by (metis leftmost_def leftmost_notword linorder_neqE_nat) lemma leftmost_Derives1: "leftmost i a \ Derives1 a j r b \ i \ j" by (metis Derives1_leftmost leftmost_unique) lemma leftmost_Derives1_propagate: assumes leftmost: "leftmost i a" and Derives1: "Derives1 a j r b" shows "(is_word b \ i = j) \ (\ k. leftmost k b \ i \ k)" proof - from leftmost_Derives1[OF leftmost Derives1] have ij: "i \ j" by auto show ?thesis proof (cases "is_word b") case True show ?thesis by (metis Derives1 True ij le_neq_implies_less leftmost leftmost_unaffected_Derives1 order_refl) next case False show ?thesis by (metis (no_types, hide_lams) Derives1 Derives1_bound Derives1_sentence2 Derives1_take append_take_drop_id ij le_neq_implies_less leftmost leftmost_append leftmost_cons_less leftmost_def length_take min.absorb2 nat_le_linear nonword_leftmost_exists not_le) qed qed lemma is_word_Derives1[elim]: "is_word a \ Derives1 a i r b \ False" by (metis Derives1_leftmost is_terminal_nonterminal is_word_is_terminal leftmost_def) lemma is_word_Derivation[elim]: "is_word a \ Derivation a D b \ D = []" by (metis Derivation_leftmost is_terminal_nonterminal is_word_def leftmost_def list_all_length) lemma leftmost_Derivation: "leftmost i a \ Derivation a D b \ j \ i \ derivation_ge D j" proof (induct D arbitrary: a b i j) case Nil thus ?case by auto next case (Cons d D) then have "\ x. Derives1 a (fst d) (snd d) x \ Derivation x D b" by auto then obtain x where ax:"Derives1 a (fst d) (snd d) x" and xb:"Derivation x D b" by blast note ji = Cons(4) note i_fstd = leftmost_Derives1[OF Cons(2) ax] note disj = leftmost_Derives1_propagate[OF Cons(2) ax] thus ?case proof(induct rule: disjCases2) case 1 with xb have "D = []" by auto with 1 ji show ?case by (simp add: derivation_ge_def) next case 2 then obtain k where k: "leftmost k x" and ik: "i \ k" by blast note ge = Cons(1)[OF k xb, where j=j] from ji ik i_fstd ge show ?case by (simp add: derivation_ge_cons) qed qed lemma derivation_ge_list_all: "derivation_ge D i = list_all (\ d. fst d \ i) D" by (simp add: Ball_set derivation_ge_def) lemma split_derivation_leftmost: assumes "derivation_ge D i" and "\ (derivation_ge D (Suc i))" shows "\ E F r. D = E@[(i, r)]@F \ derivation_ge E (Suc i)" proof - from assms have "\ k. k < length D \ fst(D ! k) \ i \ \(fst(D ! k) \ Suc i)" by (metis derivation_ge_def in_set_conv_nth) then have "\ k. k < length D \ fst(D ! k) = i" by auto then show ?thesis proof(induct rule: ex_minimal_witness) case (Minimal k) then have k_len: "k < length D" and k_i: "fst (D ! k) = i" by auto let ?E = "take k D" let ?r = "snd (D ! k)" let ?F = "drop (Suc k) D" note split = split_list_at[OF k_len] from k_i have D_k: "D ! k = (i, ?r)" by auto show ?case apply (rule exI[where x="?E"]) apply (rule exI[where x="?F"]) apply (rule exI[where x="?r"]) apply (subst D_k[symmetric]) apply (rule conjI) apply (rule split) by (metis (mono_tags, lifting) Minimal.hyps(2) Suc_leI assms(1) derivation_ge_list_all le_neq_implies_less list_all_length list_all_take) qed qed lemma Derives1_Derives1_swap: assumes "i < j" and "Derives1 a j p b" and "Derives1 b i q c" shows "\ b'. Derives1 a i q b' \ Derives1 b' (j - 1 + length (snd q)) p c" proof - from Derives1_split[OF assms(2)] obtain a1 a2 where a_src: "a = a1 @ [fst p] @ a2" and a_dest: "b = a1 @ snd p @ a2" and a1_len: "length a1 = j" by blast note a = this from a have is_sentence_a1: "is_sentence a1" using Derives1_sentence2 assms(2) is_sentence_concat by blast from a have is_sentence_a2: "is_sentence a2" using Derives1_sentence2 assms(2) is_sentence_concat by blast from a have is_symbol_fst_p: "is_symbol (fst p)" by (metis Derives1_sentence1 assms(2) is_sentence_concat is_sentence_cons) from Derives1_split[OF assms(3)] obtain b1 b2 where b_src: "b = b1 @ [fst q] @ b2" and b_dest: "c = b1 @ snd q @ b2" and b1_len: "length b1 = i" by blast have a_take_j: "a1 = take j a" by (metis a1_len a_src append_eq_conv_conj) have b_take_i: "b1 @ [fst q] = take (Suc i) b" by (metis append_assoc append_eq_conv_conj b1_len b_src length_append_singleton) from a_take_j b_take_i take_eq_take_append[where j=j and i="Suc i" and a=a] have "\ u. a1 = (b1 @ [fst q]) @ u" by (metis le_iff_add Suc_leI a1_len a_dest append_eq_conv_conj assms(1) take_add) then obtain u where u1: "a1 = (b1 @ [fst q]) @ u" by blast then have j_i_u: "j = i + 1 + length u" using Suc_eq_plus1 a1_len b1_len length_append length_append_singleton by auto from u1 is_sentence_a1 have is_sentence_b1_u: "is_sentence b1 \ is_sentence u" using is_sentence_concat by blast have u2: "b2 = u @ snd p @ a2" by (metis a_dest append_assoc b_src same_append_eq u1) let ?b = "b1 @ (snd q) @ u @ [fst p] @ a2" from assms have q_dom: "q \ \" by auto have a_b': "Derives1 a i q ?b" apply (subst Derives1_def) apply (rule exI[where x="b1"]) apply (rule exI[where x="u@[fst p]@a2"]) apply (rule exI[where x="fst q"]) apply (rule exI[where x="snd q"]) apply (auto simp add: b1_len is_sentence_cons is_sentence_concat is_sentence_a2 is_symbol_fst_p is_sentence_b1_u a_src u1 q_dom) done from assms have p_dom: "p \ \" by auto have is_sentence_snd_q: "is_sentence (snd q)" using Derives1_sentence2 a_b' is_sentence_concat by blast have b'_c: "Derives1 ?b (j - 1 + length (snd q)) p c" apply (subst Derives1_def) apply (rule exI[where x="b1 @ (snd q) @ u"]) apply (rule exI[where x="a2"]) apply (rule exI[where x="fst p"]) apply (rule exI[where x="snd p"]) apply (auto simp add: is_sentence_concat is_sentence_b1_u is_sentence_a2 p_dom is_sentence_snd_q b_dest u2 b1_len j_i_u) done show ?thesis apply (rule exI[where x="?b"]) apply (rule conjI) apply (rule a_b') apply (rule b'_c) done qed definition derivation_shift :: "derivation \ nat \ nat \ derivation" where "derivation_shift D left right = map (\ d. (fst d - left + right, snd d)) D" lemma derivation_shift_empty[simp]: "derivation_shift [] left right = []" by (auto simp add: derivation_shift_def) lemma derivation_shift_cons[simp]: "derivation_shift (d#D) left right = ((fst d - left + right, snd d)#(derivation_shift D left right))" by (simp add: derivation_shift_def) lemma Derivation_append: "Derivation a (D@E) c = (\ b. Derivation a D b \ Derivation b E c)" proof(induct D arbitrary: a c E) case Nil thus ?case by auto next case (Cons d D) thus ?case by auto qed lemma Derivation_implies_append: "Derivation a D b \ Derivation b E c \ Derivation a (D@E) c" using Derivation_append by blast lemma Derivation_swap_single_end_to_front: "i < j \ derivation_ge D j \ Derivation a (D@[(i,r)]) b \ Derivation a ((i,r)#(derivation_shift D 1 (length (snd r)))) b" proof(induct D arbitrary: a) case Nil thus ?case by auto next case (Cons d D) from Cons have "\ c. Derives1 a (fst d) (snd d) c \ Derivation c (D @ [(i, r)]) b" by simp then obtain c where ac: "Derives1 a (fst d) (snd d) c" and cb: "Derivation c (D @ [(i, r)]) b" by blast from Cons(3) have D_j: "derivation_ge D j" by (simp add: derivation_ge_cons) from Cons(1)[OF Cons(2) D_j cb, simplified] obtain x where cx: "Derives1 c i r x" and xb: "Derivation x (derivation_shift D 1 (length (snd r))) b" by auto have i_fst_d: "i < fst d" using Cons derivation_ge_cons by auto from Derives1_Derives1_swap[OF i_fst_d ac cx] obtain b' where ab': "Derives1 a i r b'" and b'x: "Derives1 b' (fst d - 1 + length (snd r)) (snd d) x" by blast show ?case using ab' b'x xb by auto qed lemma Derivation_swap_single_mid_to_front: assumes "i < j" and "derivation_ge D j" and "Derivation a (D@[(i,r)]@E) b" shows "Derivation a ((i,r)#((derivation_shift D 1 (length (snd r)))@E)) b" proof - from assms have "\ x. Derivation a (D@[(i, r)]) x \ Derivation x E b" using Derivation_append by auto then obtain x where ax: "Derivation a (D@[(i, r)]) x" and xb: "Derivation x E b" by blast with assms have "Derivation a ((i, r)#(derivation_shift D 1 (length (snd r)))) x" using Derivation_swap_single_end_to_front by blast then show ?thesis using Derivation_append xb by auto qed lemma length_derivation_shift[simp]: "length(derivation_shift D left right) = length D" by (simp add: derivation_shift_def) definition LeftDerives1 :: "sentence \ nat \ rule \ sentence \ bool" where "LeftDerives1 u i r v = (leftmost i u \ Derives1 u i r v)" lemma LeftDerives1_implies_leftderives1: "LeftDerives1 u i r v \ leftderives1 u v" by (metis Derives1_def LeftDerives1_def append_eq_conv_conj leftderives1_def leftmost_def) lemma leftmost_Derives1_leftderives: "leftmost i a \ Derives1 a i r b \ leftderives b c \ leftderives a c" using LeftDerives1_def LeftDerives1_implies_leftderives1 leftderives1_implies_leftderives leftderives_trans by blast theorem Derivation_implies_leftderives_gen: "Derivation a D (u@v) \ is_word u \ (\ w. leftderives a (u@w) \ (v = [] \ w = []) \ (\ X. is_first X v \ is_first X w))" proof (induct "length D" arbitrary: D a u v) case 0 then have "a = u@v" by auto thus ?case by (rule_tac x = v in exI, auto) next case (Suc n) from Suc have "D \ []" by auto with Suc Derivation_leftmost have "\ i. leftmost i a" by auto then obtain i where i: "leftmost i a" by blast show "?case" proof (cases "derivation_ge D (Suc i)") case True with Suc have leftmost: "leftmost i (u@v)" by (rule_tac leftmost_unaffected_Derivation[OF True i], auto) have length_u: "length u \ i" using leftmost_append[OF leftmost Suc(4)] . have take_Suc: "take (Suc i) a = take (Suc i) (u@v)" using Derivation_take[OF True Suc(3)] . with length_u have is_prefix_u: "is_prefix u a" by (metis append_assoc append_take_drop_id dual_order.strict_implies_order is_prefix_def le_imp_less_Suc take_all take_append) have a: "u @ drop (length u) a = a" using is_prefix_unsplit[OF is_prefix_u] . from take_Suc have length_take_Suc: "length (take (Suc i) a) = Suc i" by (metis Suc_leI i leftmost_def length_take min.absorb2) have v: "v \ []" proof(cases "v = []") case False thus ?thesis by auto next case True with length_u have right: "length(take (Suc i) (u@v)) = length u" by simp note left = length_take_Suc from left right take_Suc have "Suc i = length u" by auto with length_u show ?thesis by auto qed then have "\ X w. v = X#w" by (cases v, auto) then obtain X w where v: "v = X#w" by blast have is_first_X: "is_first X (drop (length u) a)" apply (rule_tac is_first_drop_length[where v=v and w=w and k="Suc i"]) apply (simp_all add: take_Suc v) apply (metis Suc_leI i leftmost_def) apply (insert length_u) apply arith done show ?thesis apply (rule exI[where x="drop (length u) a"]) by (simp add: a v is_first_cons is_first_X) next case False have Di: "derivation_ge D i" using leftmost_Derivation[OF i Suc(3), where j=i, simplified] . from split_derivation_leftmost[OF Di False] obtain E F r where D_split: "D = E @ [(i, r)] @ F" and E_Suc: "derivation_ge E (Suc i)" by auto let ?D = "(derivation_shift E 1 (length (snd r)))@F" from D_split have "Derivation a ((i,r) # ?D) (u @ v)" using Derivation_swap_single_mid_to_front E_Suc Suc.prems(1) lessI by blast then have "\ y. Derives1 a i r y \ Derivation y ?D (u @ v)" by simp then obtain y where ay:"Derives1 a i r y" and yuv: "Derivation y ?D (u @ v)" by blast have length_D': "length ?D = n" using D_split Suc.hyps(2) by auto from Suc(1)[OF length_D'[symmetric] yuv Suc(4)] obtain w where "leftderives y (u @ w)" and "(v = [] \ w = [])" and "\X. is_first X v \ is_first X w" by blast then show ?thesis using ay i leftmost_Derives1_leftderives by blast qed qed lemma derives_implies_leftderives_gen: "derives a (u@v) \ is_word u \ (\ w. leftderives a (u@w) \ (v = [] \ w = []) \ (\ X. is_first X v \ is_first X w))" using Derivation_implies_leftderives_gen derives_implies_Derivation by blast lemma derives_implies_leftderives: "derives a b \ is_word b \ leftderives a b" using derives_implies_leftderives_gen by force fun LeftDerivation :: "sentence \ derivation \ sentence \ bool" where "LeftDerivation a [] b = (a = b)" | "LeftDerivation a (d#D) b = (\ x. LeftDerives1 a (fst d) (snd d) x \ LeftDerivation x D b)" lemma LeftDerives1_implies_Derives1: "LeftDerives1 a i r b \ Derives1 a i r b" by (metis LeftDerives1_def) lemma LeftDerivation_implies_Derivation: "LeftDerivation a D b \ Derivation a D b" proof (induct D arbitrary: a) case (Nil) thus ?case by simp next case (Cons d D) thus ?case using CFG.LeftDerivation.simps(2) CFG_axioms Derivation.simps(2) LeftDerives1_implies_Derives1 by blast qed lemma LeftDerivation_implies_leftderives: "LeftDerivation a D b \ leftderives a b" proof (induct D arbitrary: a b) case Nil thus ?case by simp next case (Cons d D) note ihyps = this from ihyps have "\ x. LeftDerives1 a (fst d) (snd d) x \ LeftDerivation x D b" by auto then obtain x where "LeftDerives1 a (fst d) (snd d) x" and xb: "LeftDerivation x D b" by blast with LeftDerives1_implies_leftderives1 have d1: "leftderives a x" by auto from ihyps xb have d2:"leftderives x b" by simp show "leftderives a b" by (rule leftderives_trans[OF d1 d2]) qed lemma leftmost_witness[simp]: "leftmost (length x) (x@(N#y)) = (is_word x \ is_nonterminal N)" by (auto simp add: leftmost_def) lemma leftderives1_implies_LeftDerives1: assumes leftderives1: "leftderives1 u v" shows "\ i r. LeftDerives1 u i r v" proof - from leftderives1 have "\x y N \. u = x @ [N] @ y \ v = x @ \ @ y \ is_word x \ is_sentence y \ (N, \) \ \" by (simp add: leftderives1_def) then obtain x y N \ where u:"u = x @ [N] @ y" and v:"v = x @ \ @ y" and x:"is_word x" and y:"is_sentence y" and r:"(N, \) \ \" by blast show ?thesis apply (rule_tac x="length x" in exI) apply (rule_tac x="(N, \)" in exI) apply (auto simp add: LeftDerives1_def) apply (simp add: leftmost_def x u rule_nonterminal_type[OF r]) apply (simp add: Derives1_def u v) apply (rule_tac x=x in exI) apply (rule_tac x=y in exI) apply (auto simp add: x y r) done qed lemma LeftDerivation_LeftDerives1: "LeftDerivation a S y \ LeftDerives1 y i r z \ LeftDerivation 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 LeftDerivation.simps(2) append_Cons) qed lemma leftderives_implies_LeftDerivation: "leftderives a b \ \ D. LeftDerivation a D b" proof (induct rule: leftderives_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: "LeftDerivation a D y" by blast from ihyps leftderives1_implies_LeftDerives1 obtain i r where yz: "LeftDerives1 y i r z" by blast from LeftDerivation_LeftDerives1[OF ay yz] show ?case by auto qed lemma LeftDerivation_append: "LeftDerivation a (D@E) c = (\ b. LeftDerivation a D b \ LeftDerivation b E c)" proof(induct D arbitrary: a c E) case Nil thus ?case by auto next case (Cons d D) thus ?case by auto qed lemma LeftDerivation_implies_append: "LeftDerivation a D b \ LeftDerivation b E c \ LeftDerivation a (D@E) c" using LeftDerivation_append by blast lemma Derivation_unique_dest: "Derivation a D b \ Derivation a D c \ b = c" apply (induct D arbitrary: a b c) apply auto using Derives1_unique_dest by blast lemma Derivation_unique_src: "Derivation a D c \ Derivation b D c \ a = b" apply (induct D arbitrary: a b c) apply auto using Derives1_unique_src by blast lemma LeftDerives1_unique: "LeftDerives1 a i r b \ LeftDerives1 a j s b \ i = j \ r = s" using Derives1_def LeftDerives1_def leftmost_unique by auto lemma leftlang: "\ = { v | v. is_word v \ is_leftderivation v }" by (metis (no_types, lifting) CFG.is_derivation_def CFG.is_leftderivation_def CFG.leftderivation_implies_derivation CFG_axioms Collect_cong \_def derives_implies_leftderives) lemma leftprefixlang: "\\<^sub>P = { u | u v. is_word u \ is_leftderivation (u@v) }" apply (auto simp add: \\<^sub>P_def) using derives_implies_leftderives_gen is_derivation_def is_leftderivation_def apply blast using leftderivation_implies_derivation by blast lemma derives_implies_leftderives_cons: "is_word a \ derives u (a@X#b) \ \ c. leftderives u (a@X#c)" by (metis append_Cons append_Nil derives_implies_leftderives_gen is_first_def) lemma is_word_append[simp]: "is_word (a@b) = (is_word a \ is_word b)" by (auto simp add: is_word_def) lemma \\<^sub>P_split: "a@b \ \\<^sub>P \ a \ \\<^sub>P" by (auto simp add: \\<^sub>P_def) lemma \\<^sub>P_is_word: "a \ \\<^sub>P \ is_word a" by (metis (no_types, lifting) leftprefixlang mem_Collect_eq) definition Derive :: "sentence \ derivation \ sentence" where "Derive a D = (THE b. Derivation a D b)" lemma Derivation_dest_ex_unique: "Derivation a D b \ \! x. Derivation a D x" using CFG.Derivation_unique_dest CFG_axioms by blast lemma Derive: assumes ab: "Derivation a D b" shows "Derive a D = b" proof - note the1_equality[OF Derivation_dest_ex_unique[OF ab] ab] thus ?thesis by (simp add: Derive_def) qed end end diff --git a/thys/LocalLexing/Ladder.thy b/thys/LocalLexing/Ladder.thy --- a/thys/LocalLexing/Ladder.thy +++ b/thys/LocalLexing/Ladder.thy @@ -1,3579 +1,3579 @@ theory Ladder imports TheoremD9 begin context LocalLexing begin definition LeftDerivationFix :: "sentence \ nat \ derivation \ nat \ sentence \ bool" where "LeftDerivationFix \ i D j \ = (is_sentence \ \ is_sentence \ \ LeftDerivation \ D \ \ i < length \ \ j < length \ \ \ ! i = \ ! j \ (\ E F. D = E@(derivation_shift F 0 (Suc j)) \ LeftDerivation (take i \) E (take j \) \ LeftDerivation (drop (Suc i) \) F (drop (Suc j) \)))" definition LeftDerivationIntro :: "sentence \ nat \ rule \ nat \ derivation \ nat \ sentence \ bool" where "LeftDerivationIntro \ i r ix D j \ = (\ \. LeftDerives1 \ i r \ \ ix < length (snd r) \ (snd r) ! ix = \ ! j \ LeftDerivationFix \ (i + ix) D j \)" lemma LeftDerivationFix_empty[simp]: "is_sentence \ \ i < length \ \ LeftDerivationFix \ i [] i \" apply (auto simp add: LeftDerivationFix_def) apply (rule_tac x="[]" in exI) apply auto done lemma Derive_empty[simp]: "Derive a [] = a" by (auto simp add: Derive_def) lemma LeftDerivation_append1: "LeftDerivation a (D@[(i, r)]) c \ \ b. LeftDerivation a D b \ LeftDerives1 b i r c" by (simp add: LeftDerivation_append) lemma Derivation_append1: "Derivation a (D@[(i, r)]) c \ \ b. Derivation a D b \ Derives1 b i r c" by (simp add: Derivation_append) lemma Derivation_take_derive: assumes "Derivation a D b" shows "Derivation a (take n D) (Derive a (take n D))" by (metis Derivation_append Derive append_take_drop_id assms) lemma LeftDerivation_take_derive: assumes "LeftDerivation a D b" shows "LeftDerivation a (take n D) (Derive a (take n D))" by (metis Derive LeftDerivation_append LeftDerivation_implies_Derivation append_take_drop_id assms) lemma Derivation_Derive_take_Derives1: assumes "N \ 0" assumes "N \ length D" assumes "Derivation a D b" assumes \: "\ = Derive a (take (N - 1) D)" assumes "\ = Derive a (take N D)" shows "Derives1 \ (fst (D ! (N - 1))) (snd (D ! (N - 1))) \" proof - let ?D1 = "take (N - 1) D" let ?D2 = "take N D" from assms have app: "?D2 = ?D1 @ [D ! (N - 1)]" apply auto by (metis Suc_less_eq Suc_pred le_imp_less_Suc take_Suc_conv_app_nth) from assms have "Derivation a ?D2 \" using Derivation_take_derive by blast with app show ?thesis using Derivation.simps Derivation_append Derive \ by auto qed lemma LeftDerivation_Derive_take_LeftDerives1: assumes "N \ 0" assumes "N \ length D" assumes "LeftDerivation a D b" assumes \: "\ = Derive a (take (N - 1) D)" assumes "\ = Derive a (take N D)" shows "LeftDerives1 \ (fst (D ! (N - 1))) (snd (D ! (N - 1))) \" proof - let ?D1 = "take (N - 1) D" let ?D2 = "take N D" from assms have app: "?D2 = ?D1 @ [D ! (N - 1)]" apply auto by (metis Suc_less_eq Suc_pred le_imp_less_Suc take_Suc_conv_app_nth) from assms have "LeftDerivation a ?D2 \" using LeftDerivation_take_derive by blast with app show ?thesis by (metis Derive LeftDerivation_append1 LeftDerivation_implies_Derivation \ prod.collapse) qed lemma LeftDerives1_skip_prefix: "length a \ i \ LeftDerives1 (a@b) i r (a@c) \ LeftDerives1 b (i - length a) r c" apply (auto simp add: LeftDerives1_def) using leftmost_skip_prefix apply blast by (simp add: Derives1_skip_prefix) lemma LeftDerives1_skip_suffix: assumes i: "i < length a" assumes D: "LeftDerives1 (a@c) i r (b@c)" shows "LeftDerives1 a i r b" proof - note Derives1_def[where u="a@c" and v="b@c" and i=i and r=r] then have "\x y N \. a @ c = x @ [N] @ y \ b @ c = x @ \ @ y \ is_sentence x \ is_sentence y \ (N, \) \ \ \ r = (N, \) \ i = length x" using D LeftDerives1_implies_Derives1 by auto then obtain x y N \ where split: "a @ c = x @ [N] @ y \ b @ c = x @ \ @ y \ is_sentence x \ is_sentence y \ (N, \) \ \ \ r = (N, \) \ i = length x" by blast from split have "length (a@c) = length (x @ [N] @ y)" by auto then have "length a + length c = length x + length y + 1" by simp with split have "length a + length c = i + length y + 1" by simp with i have len_c_y: "length c \ length y" by arith let ?y = "take (length y - length c) y" from split have ac: "a @ c = (x @ [N]) @ y" by auto note cancel_suffix[where a=a and c = c and b = "x@[N]" and d = "y", OF ac len_c_y] then have a: "a = x @ [N] @ ?y" by auto from split have bc: "b @ c = (x @ \) @ y" by auto note cancel_suffix[where a=b and c = c and b = "x@\" and d = "y", OF bc len_c_y] then have b: "b = x @ \ @ ?y" by auto from split len_c_y a b show ?thesis apply (simp only: LeftDerives1_def Derives1_def) apply (rule_tac conjI) using D LeftDerives1_def i leftmost_cons_less apply blast apply (rule_tac x=x in exI) apply (rule_tac x="?y" in exI) apply (rule_tac x="N" in exI) apply (rule_tac x="\" in exI) apply auto by (rule is_sentence_take) qed lemma LeftDerives1_X_is_part_of_rule[consumes 2, case_names Suffix Prefix]: assumes aXb: "LeftDerives1 \ i r (a@[X]@b)" assumes split: "splits_at \ i \ N \" assumes prefix: "\ \. \ = a @ [X] @ \ \ length a < i \ is_word (a @ [X]) \ LeftDerives1 \ (i - length a - 1) r b \ False" assumes suffix: "\ \. \ = \ @ [X] @ b \ LeftDerives1 \ i r a \ False" shows "\ u v. a = \ @ u \ b = v @ \ \ (snd r) = u@[X]@v" proof - have aXb_old: "Derives1 \ i r (a@[X]@b)" using LeftDerives1_implies_Derives1 aXb by blast have prefix_or: "is_prefix \ a \ is_proper_prefix a \" by (metis Derives1_prefix split aXb_old is_prefix_eq_proper_prefix) have is_word_\: "is_word \" using LeftDerives1_splits_at_is_word aXb assms(2) by blast have "is_proper_prefix a \ \ False" proof - assume proper:"is_proper_prefix a \" then have "\ u. u \ [] \ \ = a@u" by (metis is_proper_prefix_def) then obtain u where u: "u \ [] \ \ = a@u" by blast note splits_at = splits_at_\[OF aXb_old split] splits_at_combine[OF split] from splits_at have \1: "\ = take i \" by blast from splits_at have \2: "\ = take i (a@[X]@b)" by blast from splits_at have len\: "length \ = i" by blast with proper have lena: "length a < i" using append_eq_conv_conj drop_eq_Nil leI u by auto with is_word_\ \2 have is_word_aX: "is_word (a@[X])" by (simp add: is_word_terminals not_less take_Cons' u) from u \2 have "a@u = take i (a@[X]@b)" by auto with lena have "u = take (i - length a) ([X]@b)" by (simp add: less_or_eq_imp_le) with lena have uX: "u = [X]@(take (i - length a - 1) b)" by (simp add: not_less take_Cons') let ?\ = "(take (i - length a - 1) b) @ [N] @ \" from splits_at have f1: "\ = \ @ [N] @ \" by blast with u uX have f2: "\ = a @ [X] @ ?\" by simp note skip = LeftDerives1_skip_prefix[where a = "a @ [X]" and b = "?\" and r = r and i = i and c = b] then have D: "LeftDerives1 ?\ (i - length a - 1) r b" using One_nat_def Suc_leI aXb append_assoc diff_diff_left f2 lena length_Cons length_append length_append_singleton list.size(3) by fastforce note prefix[OF f2 lena is_word_aX D] then show "False" . qed with prefix_or have is_prefix: "is_prefix \ a" by blast from aXb have aXb': "LeftDerives1 \ i r ((a@[X])@b)" by auto then have aXb'_old: "Derives1 \ i r ((a@[X])@b)" by (simp add: LeftDerives1_implies_Derives1) note Derives1_suffix[OF aXb'_old split] then have suffix_or: "is_suffix \ b \ is_proper_suffix b \" by (metis is_suffix_eq_proper_suffix) have "is_proper_suffix b \ \ False" proof - assume proper: "is_proper_suffix b \" then have "\ u. u \ [] \ \ = u@b" by (metis is_proper_suffix_def) then obtain u where u: "u \ [] \ \ = u@b" by blast note splits_at = splits_at_\[OF aXb_old split] splits_at_combine[OF split] from splits_at have \1: "\ = drop (Suc i) \" by blast from splits_at have \2: "\ = drop (i + length (snd r)) (a @ [X] @ b)" by blast from splits_at have len\: "length \ = length \ - i - 1" by blast with proper have lenb: "length b < length \" by (metis is_proper_suffix_length_cmp) from u \2 have "u@b = drop (i + length (snd r)) ((a @ [X]) @ b)" by auto hence "u = drop (i + length (snd r)) (a @ [X])" by (metis drop_cancel_suffix) hence uX: "u = drop (i + length (snd r)) a @ [X]" by (metis drop_keep_last u) let ?\ = "\ @ [N] @ (drop (i + length (snd r)) a)" from splits_at have f1: "\ = \ @ [N] @ \" by blast with u uX have f2: "\ = ?\ @ [X] @ b" by simp note skip = LeftDerives1_skip_suffix[where a = "?\" and c = "[X]@b" and b="a" and r = r and i = i] have f3: "i < length (\ @ [N] @ drop (i + length (snd r)) a)" proof - have f1: "1 + i + length b = length [X] + length b + i" by (metis Groups.add_ac(2) Suc_eq_plus1_left length_Cons list.size(3) list.size(4) semiring_normalization_rules(22)) have f2: "length \ - i - 1 = length ((\ @ [N] @ drop (i + length (snd r)) a) @ [X] @ b) - Suc i" by (metis f2 length_drop splits_at(1)) have "length ([]::symbol list) \ length \ - i - 1 - length b" by (metis add_diff_cancel_right' append_Nil2 append_eq_append_conv len\ length_append u) then have "length ([]::symbol list) \ length \ + length ([N] @ drop (i + length (snd r)) a) - i" using f2 f1 by (metis Suc_eq_plus1_left add_diff_cancel_right' diff_diff_left length_append) then show ?thesis by auto qed from aXb f2 have D: "LeftDerives1 (?\ @ [X] @ b) i r (a@[X]@b)" by auto note skip[OF f3 D] note suffix[OF f2 skip[OF f3 D]] then show "False" . qed with suffix_or have is_suffix: "is_suffix \ b" by blast from is_prefix have "\ u. a = \ @ u" by (auto simp add: is_prefix_def) then obtain u where u: "a = \ @ u" by blast from is_suffix have "\ v. b = v @ \" by (auto simp add: is_suffix_def) then obtain v where v: "b = v @ \" by blast from u v splits_at_combine[OF split] aXb have D:"LeftDerives1 (\@[N]@\) i r (\@(u@[X]@v)@\)" by simp from splits_at_\[OF aXb_old split] have i: "length \ = i" by blast from i have i1: "length \ \ i" and i2: "i \ length \" by auto note LeftDerives1_skip_suffix[OF _ LeftDerives1_skip_prefix[OF i1 D], simplified, OF i2] then have "LeftDerives1 [N] 0 r (u @ [X] @ v)" by auto then have "Derives1 [N] 0 r (u @ [X] @ v)" using LeftDerives1_implies_Derives1 by auto then have r: "snd r = u @ [X] @ v" by (metis Derives1_split append_Cons append_Nil length_0_conv list.inject self_append_conv) show ?thesis using u v r by auto qed lemma LeftDerivationFix_grow_suffix: assumes LDF: "LeftDerivationFix (b1@[X]@b2) (length b1) D j c" assumes suffix_b2: "LeftDerives1 suffix e r b2" assumes is_word_b1X: "is_word (b1@[X])" shows "LeftDerivationFix (b1@[X]@suffix) (length b1) ((e + length (b1@[X]), r)#D) j c" proof - from LDF have LDF': "is_sentence (b1@[X]@b2) \ is_sentence c \ LeftDerivation (b1 @ [X] @ b2) D c \ length b1 < length (b1 @ [X] @ b2) \ j < length c \ (b1 @ [X] @ b2) ! length b1 = c ! j \ (\E F. D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take (length b1) (b1 @ [X] @ b2)) E (take j c) \ LeftDerivation (drop (Suc (length b1)) (b1 @ [X] @ b2)) F (drop (Suc j) c))" using LeftDerivationFix_def by blast then obtain E F where EF: "D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take (length b1) (b1 @ [X] @ b2)) E (take j c) \ LeftDerivation (drop (Suc (length b1)) (b1 @ [X] @ b2)) F (drop (Suc j) c)" by blast then have LD_b1_c: "LeftDerivation b1 E (take j c)" by simp with is_word_b1X have E: "E = []" using LeftDerivation_implies_Derivation is_word_Derivation is_word_append by blast then have b1_def: "b1 = take j c" using LD_b1_c by auto then have b1_len: "j = length b1" by (simp add: LDF' dual_order.strict_implies_order min.absorb2) have D: "D = derivation_shift F 0 (Suc j)" using EF E by simp have step: "LeftDerives1 (b1 @ [X] @ suffix) (Suc (e + length b1)) r (b1 @ [X] @ b2) \ LeftDerivation (b1 @ [X] @ b2) D c" by (metis LDF' LeftDerives1_append_prefix add_Suc_right append_assoc assms(2) is_word_b1X length_append_singleton) then have is_sentence_b1Xsuffix: "is_sentence (b1 @ [X] @ suffix)" using Derives1_sentence1 LeftDerives1_implies_Derives1 by blast have X_eq_cj: "X = c ! j" using LDF' by auto show ?thesis apply (simp add: LeftDerivationFix_def) apply (rule conjI) using is_sentence_b1Xsuffix apply simp apply (rule conjI) using LDF' apply simp apply (rule conjI) using step apply force apply (rule conjI) using LDF' apply simp apply (rule conjI) apply (rule X_eq_cj) apply (rule_tac x="[]" in exI) apply (rule_tac x="(e, r)#F" in exI) apply auto apply (rule b1_len[symmetric]) apply (rule D) apply (rule b1_def) apply (rule_tac x="b2" in exI) apply (simp add: suffix_b2) using EF by auto qed lemma Derives1_append_suffix: assumes Derives1: "Derives1 v i r w" assumes u: "is_sentence u" shows "Derives1 (v@u) i r (w@u)" proof - have "\ \ N \. splits_at v i \ N \" using assms splits_at_ex by auto then obtain \ N \ where split_v: "splits_at v i \ N \" by blast have split_w: "w = \@(snd r)@\" using assms split_v splits_at_combine_dest by blast have split_uv: "splits_at (v@u) i \ N (\@u)" by (simp add: split_v splits_at_append) have is_sentence_uv: "is_sentence (v@u)" using Derives1 Derives1_sentence1 is_sentence_concat u by blast show ?thesis by (metis Derives1 Derives1_nonterminal Derives1_rule append_assoc is_sentence_uv split_uv split_v split_w splits_at_implies_Derives1) qed lemma leftmost_append_suffix: "leftmost i v \ leftmost i (v@u)" by (simp add: leftmost_def nth_append) lemma LeftDerives1_append_suffix: assumes Derives1: "LeftDerives1 v i r w" assumes u: "is_sentence u" shows "LeftDerives1 (v@u) i r (w@u)" proof - have 1: "Derives1 v i r w" by (simp add: Derives1 LeftDerives1_implies_Derives1) have 2: "leftmost i v" using Derives1 LeftDerives1_def by blast have 3: "is_sentence u" using u by fastforce have 4: "Derives1 (v@u) i r (w@u)" by (simp add: "1" "3" Derives1_append_suffix) have 5: "leftmost i (v@u)" by (simp add: "2" leftmost_append_suffix u) show ?thesis by (simp add: "4" "5" LeftDerives1_def) qed lemma LeftDerivationFix_is_sentence: "LeftDerivationFix a i D j b \ is_sentence a \ is_sentence b" using LeftDerivationFix_def by blast lemma LeftDerivationIntro_is_sentence: "LeftDerivationIntro \ i r ix D j \ \ is_sentence \ \ is_sentence \" by (meson Derives1_sentence1 LeftDerivationFix_is_sentence LeftDerivationIntro_def LeftDerives1_implies_Derives1) lemma LeftDerivationFix_grow_prefix: assumes LDF: "LeftDerivationFix (b1@[X]@b2) (length b1) D j c" assumes prefix_b1: "LeftDerives1 prefix e r b1" shows "LeftDerivationFix (prefix@[X]@b2) (length prefix) ((e, r)#D) j c" proof - from LDF have LDF': "LeftDerivation (b1 @ [X] @ b2) D c \ length b1 < length (b1 @ [X] @ b2) \ j < length c \ (b1 @ [X] @ b2) ! length b1 = c ! j \ (\E F. D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take (length b1) (b1 @ [X] @ b2)) E (take j c) \ LeftDerivation (drop (Suc (length b1)) (b1 @ [X] @ b2)) F (drop (Suc j) c))" using LeftDerivationFix_def by blast then obtain E F where EF: "D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take (length b1) (b1 @ [X] @ b2)) E (take j c) \ LeftDerivation (drop (Suc (length b1)) (b1 @ [X] @ b2)) F (drop (Suc j) c)" by blast then have E_b1_c: "LeftDerivation b1 E (take j c)" by simp with EF have F_b2_c: "LeftDerivation b2 F (drop (Suc j) c)" by simp have step: "LeftDerives1 (prefix @ [X] @ b2) e r (b1 @ [X] @ b2)" using LDF LeftDerivationFix_is_sentence LeftDerives1_append_suffix is_sentence_concat prefix_b1 by blast show ?thesis apply (simp add: LeftDerivationFix_def) apply (rule conjI) apply (metis Derives1_sentence1 LDF LeftDerivationFix_def LeftDerives1_implies_Derives1 is_sentence_concat is_sentence_cons prefix_b1) apply (rule conjI) using LDF LeftDerivationFix_is_sentence apply blast apply (rule conjI) apply (rule_tac x="b1@[X]@b2" in exI) using step apply simp using LDF' apply auto[1] apply (rule conjI) using LDF' apply simp apply (rule conjI) using LDF' apply auto[1] apply (rule_tac x="(e,r)#E" in exI) apply (rule_tac x="F" in exI) apply (auto simp add: EF F_b2_c) apply (rule_tac x="b1" in exI) apply (simp add: prefix_b1 E_b1_c) done qed lemma LeftDerivationFixOrIntro: "LeftDerivation a D \ \ is_sentence \ \ j < length \ \ (\ i. LeftDerivationFix a i D j \) \ (\ d \ ix. d < length D \ LeftDerivation a (take d D) \ \ LeftDerivationIntro \ (fst (D ! d)) (snd (D ! d)) ix (drop (Suc d) D) j \)" proof (induct "length D" arbitrary: a D \ j rule: less_induct) (* The induction here is unnecessary, but we use it anyway for context reasons *) case less have "length D = 0 \ length D \ 0" by blast then show ?case proof (induct rule: disjCases2) case 1 then have D: "D = []" by auto with less have "\i. LeftDerivationFix a i D j \" apply (rule_tac x=j in exI) by auto then show ?case by blast next case 2 note less2 = 2 have "\ n \ i. n \ length D \ \ = Derive a (take n D) \ LeftDerivationFix \ i (drop n D) j \" apply (rule_tac x="length D" in exI) apply auto using Derive LeftDerivationFix_empty LeftDerivation_implies_Derivation less by blast then show ?case proof (induct rule: ex_minimal_witness) case (Minimal N) then obtain \ i where Minimal_N: "N \ length D \ \ = Derive a (take N D) \ LeftDerivationFix \ i (drop N D) j \" by blast have "N = 0 \ N \ 0" by blast then show ?case proof (induct rule: disjCases2) case 1 with Minimal_N have "\ = a" by auto with 1 Minimal_N show ?case apply (rule_tac disjI1) by auto next case 2 let ?\ = "Derive a (take (N - 1) D)" have LeftDerives1_\: "LeftDerives1 ?\ (fst (D ! (N - 1))) (snd (D ! (N - 1))) \" using "2.hyps" LeftDerivation_Derive_take_LeftDerives1 Minimal_N less.prems(1) by blast then have Derives1_\: "Derives1 ?\ (fst (D ! (N - 1))) (snd (D ! (N - 1))) \" using LeftDerives1_implies_Derives1 by blast have i_len: "i < length \" using Minimal_N by (auto simp add: LeftDerivationFix_def) then have "\ X \_1 \_2. splits_at \ i \_1 X \_2" using splits_at_def by blast then obtain X \_1 \_2 where \_split: "splits_at \ i \_1 X \_2" by blast then have \_combine: "\ = \_1 @ [X] @ \_2" using splits_at_combine by blast then have LeftDerives1_\_hyp: "LeftDerives1 ?\ (fst (D ! (N - 1))) (snd (D ! (N - 1))) (\_1 @ [X] @ \_2)" using LeftDerives1_\ by blast from \_split have i_def: "i = length \_1" by (simp add: dual_order.strict_implies_order min.absorb2 splits_at_def) have "\ Y \_1 \_2. splits_at ?\ (fst (D ! (N - 1))) \_1 Y \_2" using Derives1_\ splits_at_ex by blast then obtain Y \_1 \_2 where \_split: "splits_at ?\ (fst (D ! (N - 1))) \_1 Y \_2" by blast have NFix: "LeftDerivationFix (\_1 @ [X] @ \_2) (length \_1) (drop N D) j \" using Minimal_N \_combine i_def by auto from LeftDerives1_\_hyp \_split have "\u v. \_1 = \_1 @ u \ \_2 = v @ \_2 \ snd (snd (D ! (N - 1))) = u @ [X] @ v" proof (induct rule: LeftDerives1_X_is_part_of_rule) case (Suffix suffix) let ?k = "N - 1" let ?\ = "Derive a (take ?k D)" let ?i = "length \_1" have k_less: "?k < length D" using "2.hyps" Minimal_N by linarith then have k_leq: "?k \ length D" by auto have drop_k_d: "drop ?k D = (D ! (N - 1))#(drop N D)" using "2.hyps" Cons_nth_drop_Suc k_less by fastforce from LeftDerivationFix_grow_suffix[OF NFix Suffix(4) Suffix(3)] Suffix(1) Suffix(2) 2 have "LeftDerivationFix ?\ ?i (drop ?k D) j \" apply auto by (metis One_nat_def drop_k_d) with Minimal(2)[where k="?k"] show "False" using "2.hyps" k_leq by auto next case (Prefix prefix) have collapse: "(fst (D ! (N - 1)), snd (D ! (N - 1))) # drop N D = drop (N - 1) D" by (metis "2.hyps" Cons_nth_drop_Suc Minimal_N Suc_diff_1 neq0_conv not_less not_less_eq prod.collapse) from LeftDerivationFix_grow_prefix[OF NFix Prefix(2)] Prefix(1) collapse have "LeftDerivationFix ?\ (length prefix) (drop (N - 1) D) j \" by auto with Minimal(2)[where k = "N - 1"] show "False" by (metis Minimal_N collapse diff_le_self le_neq_implies_less less_imp_diff_less less_or_eq_imp_le not_Cons_self2) qed then obtain u v where uv: "\_1 = \_1 @ u \ \_2 = v @ \_2 \ snd (snd (D ! (N - 1))) = u @ [X] @ v" by blast have X_1: "snd (snd (D ! (N - Suc 0))) ! length u = X" using uv by auto have X_2: "\ ! j = X" using LeftDerivationFix_def NFix by auto show ?case apply (rule disjI2) apply (rule_tac x="N - 1" in exI) apply (rule_tac x="?\" in exI) apply (rule_tac x="length u" in exI) apply (rule conjI) using Minimal_N less2 apply linarith apply (rule conjI) using LeftDerivation_take_derive less.prems(1) apply blast apply (subst LeftDerivationIntro_def) apply (rule_tac x=\ in exI) apply auto using LeftDerives1_\ One_nat_def apply presburger using uv apply auto[1] using X_1 X_2 apply auto[1] by (metis (no_types, lifting) "2.hyps" Derives1_\ Derives1_split Minimal_N One_nat_def Suc_diff_1 \_split append_eq_conv_conj i_def length_append neq0_conv splits_at_def uv) qed qed qed qed type_synonym deriv = "nat \ nat \ nat" type_synonym ladder = "deriv list" definition deriv_n :: "deriv \ nat" where "deriv_n d = fst d" definition deriv_j :: "deriv \ nat" where "deriv_j d = fst (snd d)" definition deriv_ix :: "deriv \ nat" where "deriv_ix d = snd (snd d)" definition deriv_i :: "deriv \ nat" where "deriv_i d = snd (snd d)" definition ladder_j :: "ladder \ nat \ nat" where "ladder_j L index = deriv_j (L ! index)" definition ladder_i :: "ladder \ nat \ nat" where "ladder_i L index = (if index = 0 then deriv_i (hd L) else ladder_j L (index - 1))" definition ladder_n :: "ladder \ nat \ nat" where "ladder_n L index = deriv_n (L ! index)" definition ladder_prev_n :: "ladder \ nat \ nat" where "ladder_prev_n L index = (if index = 0 then 0 else (ladder_n L (index - 1)))" definition ladder_ix :: "ladder \ nat \ nat" where "ladder_ix L index = (if index = 0 then undefined else deriv_ix (L ! index))" definition ladder_last_j :: "ladder \ nat" where "ladder_last_j L = ladder_j L (length L - 1)" definition ladder_last_n :: "ladder \ nat" where "ladder_last_n L = ladder_n L (length L - 1)" definition is_ladder :: "derivation \ ladder \ bool" where "is_ladder D L = (L \ [] \ (\ u. u < length L \ ladder_n L u \ length D) \ (\ u v. u < v \ v < length L \ ladder_n L u < ladder_n L v) \ ladder_last_n L = length D)" definition ladder_\ :: "sentence \ derivation \ ladder \ nat \ sentence" where "ladder_\ a D L index = Derive a (take (ladder_n L index) D)" definition ladder_\ :: "sentence \ derivation \ ladder \ nat \ sentence" where "ladder_\ a D L index = (if index = 0 then a else ladder_\ a D L (index - 1))" definition LeftDerivationIntrosAt :: "sentence \ derivation \ ladder \ nat \ bool" where "LeftDerivationIntrosAt a D L index = ( let \ = ladder_\ a D L index in let i = ladder_i L index in let j = ladder_j L index in let ix = ladder_ix L index in let \ = ladder_\ a D L index in let n = ladder_n L (index - 1) in let m = ladder_n L index in let e = D ! n in let E = drop (Suc n) (take m D) in i = fst e \ LeftDerivationIntro \ i (snd e) ix E j \)" definition LeftDerivationIntros :: "sentence \ derivation \ ladder \ bool" where "LeftDerivationIntros a D L = ( \ index. 1 \ index \ index < length L \ LeftDerivationIntrosAt a D L index)" definition LeftDerivationLadder :: "sentence \ derivation \ ladder \ sentence \ bool" where "LeftDerivationLadder a D L b = ( LeftDerivation a D b \ is_ladder D L \ LeftDerivationFix a (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ a D L 0) \ LeftDerivationIntros a D L)" definition mk_deriv_fix :: "nat \ nat \ nat \ deriv" where "mk_deriv_fix i n j = (n, j, i)" definition mk_deriv_intro :: "nat \ nat \ nat \ deriv" where "mk_deriv_intro ix n j = (n, j, ix)" lemma mk_deriv_fix_i[simp]: "deriv_i (mk_deriv_fix i n j) = i" by (simp add: deriv_i_def mk_deriv_fix_def) lemma mk_deriv_fix_j[simp]: "deriv_j (mk_deriv_fix i n j) = j" by (simp add: deriv_j_def mk_deriv_fix_def) lemma mk_deriv_fix_n[simp]: "deriv_n (mk_deriv_fix i n j) = n" by (simp add: deriv_n_def mk_deriv_fix_def) lemma mk_deriv_intro_i[simp]: "deriv_i (mk_deriv_intro i n j) = i" by (simp add: deriv_i_def mk_deriv_intro_def) lemma mk_deriv_intro_ix[simp]: "deriv_ix (mk_deriv_intro ix n j) = ix" by (simp add: deriv_ix_def mk_deriv_intro_def) lemma mk_deriv_intro_j[simp]: "deriv_j (mk_deriv_intro i n j) = j" by (simp add: deriv_j_def mk_deriv_intro_def) lemma mk_deriv_intro_n[simp]: "deriv_n (mk_deriv_intro i n j) = n" by (simp add: deriv_n_def mk_deriv_intro_def) lemma LeftDerivationFix_implies_ex_ladder: "LeftDerivationFix a i D j \ \ \ L. LeftDerivationLadder a D L \ \ ladder_last_j L = j \ ladder_last_n L = length D" apply (rule_tac x="[mk_deriv_fix i (length D) j]" in exI) apply (auto simp add: LeftDerivationLadder_def) apply (simp add: LeftDerivationFix_def) apply (simp add: is_ladder_def) apply (auto simp add: ladder_i_def ladder_j_def ladder_n_def ladder_\_def) apply (simp add: ladder_last_n_def ladder_n_def) using Derive LeftDerivationFix_def LeftDerivation_implies_Derivation apply blast apply (simp add: LeftDerivationIntros_def) apply (simp add: ladder_last_j_def ladder_j_def) apply (simp add: ladder_last_n_def ladder_n_def) done lemma trivP[case_names prems]: "P \ P" by blast lemma LeftDerivationLadder_ladder_n_bound: assumes "LeftDerivationLadder a D L b" assumes "index < length L" shows "ladder_n L index \ length D" using LeftDerivationLadder_def assms(1) assms(2) is_ladder_def by blast lemma LeftDerivationLadder_deriv_n_bound: assumes "LeftDerivationLadder a D L b" assumes "index < length L" shows "deriv_n (L ! index) \ length D" using LeftDerivationLadder_def assms(1) assms(2) is_ladder_def ladder_n_def by auto lemma ladder_n_simp1[simp]: "u < length L \ ladder_n (L@L') u = ladder_n L u" by (simp add: ladder_n_def) lemma ladder_n_simp2[simp]: "ladder_n (L@[d]) (length L) = deriv_n d" by (simp add: ladder_n_def) lemma ladder_j_simp1[simp]: "u < length L \ ladder_j (L@L') u = ladder_j L u" by (simp add: ladder_j_def) lemma ladder_j_simp2[simp]: "ladder_j (L@[d]) (length L) = deriv_j d" by (simp add: ladder_j_def) lemma ladder_i_simp1[simp]: "u < length L \ ladder_i (L@L') u = ladder_i L u" by (auto simp add: ladder_i_def) lemma ladder_ix_simp1[simp]: "u < length L \ ladder_ix (L@L') u = ladder_ix L u" by (auto simp add: ladder_ix_def) lemma ladder_ix_simp2[simp]: "L \ [] \ ladder_ix (L@[d]) (length L) = deriv_ix d" by (auto simp add: ladder_ix_def) lemma ladder_\_simp1[simp]: "u < length L \ ladder_\ a D (L@L') u = ladder_\ a D L u" by (simp add: ladder_\_def) lemma ladder_\_simp2[simp]: "u < length L \ is_ladder D L \ ladder_\ a (D@D') L u = ladder_\ a D L u" by (simp add: is_ladder_def ladder_\_def) lemma ladder_\_simp1[simp]: "u < length L \ ladder_\ a D (L@L') u = ladder_\ a D L u" by (simp add: ladder_\_def) lemma ladder_\_simp2[simp]: "u < length L \ is_ladder D L \ ladder_\ a (D@D') L u = ladder_\ a D L u" by (simp add: is_ladder_def ladder_\_def) lemma ladder_n_minus_1_bound: "is_ladder D L \ index \ 1 \ index < length L \ ladder_n L (index - Suc 0) < length D" by (metis (no_types, lifting) One_nat_def Suc_diff_1 Suc_le_lessD dual_order.strict_implies_order is_ladder_def le_neq_implies_less not_less) lemma LeftDerivationIntrosAt_ignore_appendix: assumes is_ladder: "is_ladder D L" assumes hyp: "LeftDerivationIntrosAt a D L index" assumes index_ge: "index \ 1" assumes index_less: "index < length L" shows "LeftDerivationIntrosAt a (D @ D') (L @ L') index" proof - have index_minus_1: "index - Suc 0 < length L" using index_less by arith have is_0: "ladder_n L index - length D = 0" using index_less is_ladder is_ladder_def by auto from index_ge index_less show ?thesis apply (simp add: LeftDerivationIntrosAt_def Let_def) apply (simp add: index_minus_1 is_ladder ladder_n_minus_1_bound is_0) using hyp apply (auto simp add: LeftDerivationIntrosAt_def Let_def) done qed lemma ladder_i_eq_last_j: "L \ [] \ ladder_i (L @ L') (length L) = ladder_last_j L" by (simp add: ladder_i_def ladder_last_j_def) lemma ladder_last_n_intro: "L \ [] \ ladder_n L (length L - Suc 0) = ladder_last_n L" by (simp add: ladder_last_n_def) lemma is_ladder_not_empty: "is_ladder D L \ L \ []" using is_ladder_def by blast lemma last_ladder_\: assumes is_ladder: "is_ladder D L" assumes ladder_last_n: "ladder_last_n L = length D" shows "ladder_\ a D L (length L - Suc 0) = Derive a D" proof - from is_ladder is_ladder_not_empty have "L \ []" by blast then show ?thesis by (simp add: ladder_\_def ladder_last_n_intro ladder_last_n) qed lemma ladder_\_full: assumes is_ladder: "is_ladder D L" assumes ladder_last_n: "ladder_last_n L = length D" shows "ladder_\ a (D @ D') (L @ L') (length L) = Derive a D" proof - from is_ladder have L_not_empty: "L \ []" by (simp add: is_ladder_def) with is_ladder ladder_last_n show ?thesis apply (simp add: ladder_\_def) apply (simp add: last_ladder_\) done qed lemma LeftDerivationIntro_implies_LeftDerivation: "LeftDerivationIntro \ i r ix D j \ \ LeftDerivation \ ((i,r)#D) \" using LeftDerivationFix_def LeftDerivationIntro_def by auto lemma LeftDerivationLadder_grow: "LeftDerivationLadder a D L \ \ ladder_last_j L = i \ LeftDerivationIntro \ i r ix E j \ \ LeftDerivationLadder a (D@[(i, r)]@E) (L@[mk_deriv_intro ix (Suc(length D + length E)) j]) \" proof (induct arbitrary: a D L \ i r ix E j \ rule: trivP) case prems { fix u :: nat assume "u < Suc (length L)" then have "u < length L \ u = length L" by arith then have "ladder_n (L @ [mk_deriv_intro ix (Suc (length D + length E)) j]) u \ Suc (length D + length E)" proof (induct rule: disjCases2) case 1 then show ?case apply simp by (meson LeftDerivationLadder_ladder_n_bound le_Suc_eq le_add1 le_trans prems(1)) next case 2 then show ?case by (simp add: ladder_n_def) qed } note ladder_n_ineqs = this { fix u :: nat fix v :: nat assume u_less_v: "u < v" assume "v < Suc (length L)" then have "v < length L \ v = length L" by arith then have "ladder_n (L @ [mk_deriv_intro ix (Suc (length D + length E)) j]) u < ladder_n (L @ [mk_deriv_intro ix (Suc (length D + length E)) j]) v" proof (induct rule: disjCases2) case 1 with u_less_v have u_bound: "u < length L" by arith show ?case using 1 u_bound apply simp using prems u_less_v LeftDerivationLadder_def is_ladder_def by auto next case 2 with u_less_v have u_bound: "u < length L" by arith have "deriv_n (L ! u) \ length D" using LeftDerivationLadder_deriv_n_bound prems(1) u_bound by blast then show ?case apply (simp add: u_bound) apply (simp add: ladder_n_def 2) done qed } note ladder_n_ineqs = ladder_n_ineqs this have is_ladder: "is_ladder (D @ (i, r) # E) (L @ [mk_deriv_intro ix (Suc (length D + length E)) j])" apply (auto simp add: is_ladder_def) using ladder_n_ineqs apply auto apply (simp add: ladder_last_n_def) done have is_ladder_L: "is_ladder D L" using LeftDerivationLadder_def prems.prems(1) by blast have ladder_last_n_eq_length: "ladder_last_n L = length D" using is_ladder_L is_ladder_def by blast have L_not_empty: "L \ []" using LeftDerivationLadder_def is_ladder_def prems(1) by blast { fix index :: nat assume index_ge: "Suc 0 \ index" assume "index < Suc (length L)" then have "index < length L \ index = length L" by arith then have "LeftDerivationIntrosAt a (D @ (i, r) # E) (L @ [mk_deriv_intro ix (Suc (length D + length E)) j]) index" proof (induct rule: disjCases2) case 1 then show ?case using LeftDerivationIntrosAt_ignore_appendix LeftDerivationIntros_def LeftDerivationLadder_def One_nat_def index_ge prems.prems(1) by presburger next case 2 have min_simp: "\ n E. min n (Suc (n + length E)) = n" by auto with 2 prems is_ladder_L ladder_last_n_eq_length show ?case apply (simp add: LeftDerivationIntrosAt_def) apply (simp add: L_not_empty ladder_i_eq_last_j ladder_last_n_intro) apply (simp add: ladder_\_full min_simp) apply (simp add: ladder_\_def) by (metis Derive LeftDerivationIntro_implies_LeftDerivation LeftDerivationLadder_def LeftDerivation_implies_Derivation LeftDerivation_implies_append) qed } then show ?case apply (auto simp add: LeftDerivationLadder_def) using prems apply (auto simp add: LeftDerivationLadder_def)[1] using LeftDerivationFix_def LeftDerivationIntro_def LeftDerivation_append apply auto[1] using is_ladder apply simp using L_not_empty apply simp using LeftDerivationLadder_def LeftDerivationLadder_ladder_n_bound ladder_\_def prems.prems(1) apply auto[1] apply (subst LeftDerivationIntros_def) apply auto done qed lemma LeftDerivationIntro_bounds_ij: "LeftDerivationIntro \ i r ix D j \ \ i < length \ \ j < length \" by (meson Derives1_bound LeftDerivationFix_def LeftDerivationIntro_def LeftDerives1_implies_Derives1) theorem LeftDerivationLadder_exists: "LeftDerivation a D \ \ is_sentence \ \ j < length \ \ \ L. LeftDerivationLadder a D L \ \ ladder_last_j L = j" proof (induct "length D" arbitrary: a D \ j rule: less_induct) case less from LeftDerivationFixOrIntro[OF less(2,3,4)] show ?case proof (induct rule: disjCases2) case 1 then obtain i where "LeftDerivationFix a i D j \" by blast show ?case using "1.hyps" LeftDerivationFix_implies_ex_ladder by blast next case 2 then obtain d \ ix where inductrule: "d < length D \ LeftDerivation a (take d D) \ \ LeftDerivationIntro \ (fst (D ! d)) (snd (D ! d)) ix (drop (Suc d) D) j \" by blast then have less_length_D: "length (take d D) < length D" and LeftDerivation_\: "LeftDerivation a (take d D) \" by auto have is_sentence_\: "is_sentence \" using LeftDerivationIntro_is_sentence inductrule by blast have "fst (D ! d) < length \" using LeftDerivationIntro_bounds_ij inductrule by blast from less(1)[OF less_length_D LeftDerivation_\ is_sentence_\, where j=" fst (D ! d)", OF this] obtain L where induct_Ladder: "LeftDerivationLadder a (take d D) L \" and induct_last: "ladder_last_j L = fst (D ! d)" by blast have induct_intro: "LeftDerivationIntro \ (fst (D ! d)) (snd (D ! d)) ix (drop (Suc d) D) j \" using inductrule by blast have "d < length D" using inductrule by blast then have simp_to_D: "take d D @ D ! d # drop (Suc d) D = D" using id_take_nth_drop by force from LeftDerivationLadder_grow[OF induct_Ladder induct_last induct_intro] simp_to_D show ?case apply auto apply (rule_tac x= "L @ [mk_deriv_intro ix (Suc (min (length D) d + (length D - Suc d))) j]" in exI) apply (simp add: ladder_last_j_def) done qed qed lemma LeftDerivationLadder_L_0: assumes "LeftDerivationLadder \ D L \" assumes "length L = 1" shows "\ i. LeftDerivationFix \ i D (ladder_last_j L) \" proof - have "is_ladder D L" using assms by (auto simp add: LeftDerivationLadder_def) then have ladder_n: "ladder_n L 0 = length D" by (simp add: assms(2) is_ladder_def ladder_last_n_def) show ?thesis apply (rule_tac x = "ladder_i L 0" in exI) using assms(1) apply (auto simp add: LeftDerivationLadder_def) by (metis Derive LeftDerivationFix_def LeftDerivation_implies_Derivation One_nat_def assms(2) diff_Suc_1 ladder_last_j_def ladder_n order_refl take_all) qed lemma LeftDerivationFix_splits_at_derives: assumes "LeftDerivationFix a i D j b" shows "\ U a1 a2 b1 b2. splits_at a i a1 U a2 \ splits_at b j b1 U b2 \ derives a1 b1 \ derives a2 b2" proof - note hyp = LeftDerivationFix_def[where \=a and \=b and D=D and i=i and j=j] from hyp obtain E F where EF: "D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take i a) E (take j b) \ LeftDerivation (drop (Suc i) a) F (drop (Suc j) b)" using assms by blast show ?thesis apply (rule_tac x="a ! i" in exI) apply (rule_tac x="take i a" in exI) apply (rule_tac x="drop (Suc i) a" in exI) apply (rule_tac x="take j b" in exI) apply (rule_tac x="drop (Suc j) b" in exI) using Derivation_implies_derives LeftDerivation_implies_Derivation assms hyp splits_at_def by blast qed lemma LeftDerivation_append_suffix: "LeftDerivation a D b \ is_sentence c \ LeftDerivation (a@c) D (b@c)" proof (induct D arbitrary: a b c) case Nil then show ?case by auto next case (Cons d D) then show ?case apply auto apply (rule_tac x="x@c" in exI) apply auto using LeftDerives1_append_suffix by simp qed lemma LeftDerivation_impossible: "LeftDerivation a D b \ i < length a \ is_nonterminal (a ! i) \ derivation_ge D (Suc i) \ D = []" proof (induct D) case Nil then show ?case by auto next case (Cons d D) then have lm: "\ j. leftmost j a \ j \ i" by (metis Derives1_sentence1 LeftDerivation.simps(2) LeftDerives1_implies_Derives1 leftmost_exists leftmost_unique) from Cons show ?case apply auto apply (auto simp add: derivation_ge_def LeftDerives1_def) using lm[where j="fst d"] by arith qed lemma derivation_ge_shift: "derivation_ge (derivation_shift F 0 j) j" apply (induct F) apply (auto simp add: derivation_ge_def) done lemma LeftDerivationFix_splits_at_nonterminal: assumes "LeftDerivationFix a i D j b" assumes "is_nonterminal (a ! i)" shows "\ U a1 a2 b1. splits_at a i a1 U a2 \ splits_at b j b1 U a2 \ LeftDerivation a1 D b1" proof - note hyp = LeftDerivationFix_def[where \=a and \=b and D=D and i=i and j=j] from hyp obtain E F where EF: "D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take i a) E (take j b) \ LeftDerivation (drop (Suc i) a) F (drop (Suc j) b)" using assms by blast have "\ \. LeftDerivation a E \ \ LeftDerivation \ (derivation_shift F 0 (Suc j)) b" using EF LeftDerivation_append assms(1) hyp by blast then obtain \ where \_intro: "LeftDerivation a E \ \ LeftDerivation \ (derivation_shift F 0 (Suc j)) b" by blast have "LeftDerivation ((take i a)@(drop i a)) E ((take j b)@(drop i a))" by (metis EF LeftDerivation_append_suffix append_take_drop_id assms(1) hyp is_sentence_concat) then have "LeftDerivation a E ((take j b)@(drop i a))" by simp then have \_decomposed: "\ = (take j b)@(drop i a)" using Derivation_unique_dest LeftDerivation_implies_Derivation \_intro by blast then have "\ ! j = a ! i" by (metis Cons_nth_drop_Suc assms(1) hyp length_take min.absorb2 nth_append_length order.strict_implies_order) then have is_nt: "is_nonterminal (\ ! j)" by (simp add: assms(2)) have index_j: "j < length \" using \_decomposed assms(1) hyp by auto have derivation: "LeftDerivation \ (derivation_shift F 0 (Suc j)) b" by (simp add: \_intro) from LeftDerivation_impossible[OF derivation index_j is_nt derivation_ge_shift] have F: "F = []" by (metis length_0_conv length_derivation_shift) then have \_is_b: "\ = b" using \_intro by auto show ?thesis apply (rule_tac x="a ! i" in exI) apply (rule_tac x="take i a" in exI) apply (rule_tac x="drop (Suc i) a" in exI) apply (rule_tac x="take j b" in exI) using EF F assms(1) hyp splits_at_def by auto qed lemma LeftDerivationIntro_implies_nonterminal: "LeftDerivationIntro \ i (snd e) ix E j \ \ is_nonterminal (\ ! i)" by (simp add: LeftDerivationIntro_def LeftDerives1_def leftmost_is_nonterminal) lemma LeftDerivationIntrosAt_implies_nonterminal: "LeftDerivationIntrosAt a D L index \ is_nonterminal((ladder_\ a D L index) ! (ladder_i L index))" by (meson LeftDerivationIntro_implies_nonterminal LeftDerivationIntrosAt_def) lemma LeftDerivationIntro_examine_rule: "LeftDerivationIntro \ i r ix D j \ \ splits_at \ i \1 M \2 \ \ \. M = fst r \ \ = snd r \ (M, \) \ \" by (metis Derives1_nonterminal Derives1_rule LeftDerivationIntro_def LeftDerives1_implies_Derives1 prod.collapse) lemma LeftDerivation_skip_prefixword_ex: assumes "LeftDerivation (u@v) D w" assumes "is_word u" shows "\ w'. w = u@w' \ LeftDerivation v (derivation_shift D (length u) 0) w'" by (metis LeftDerivation.simps(1) LeftDerivation_breakdown LeftDerivation_implies_Derivation LeftDerivation_skip_prefix append_eq_conv_conj assms(1) assms(2) is_word_Derivation is_word_Derivation_derivation_ge) definition ladder_cut :: "ladder \ nat \ ladder" where "ladder_cut L n = (let i = length L - 1 in L[i := (n, snd (L ! i))])" fun deriv_shift :: "nat \ nat \ deriv \ deriv" where "deriv_shift dn dj (n, j, i) = (n - dn, j - dj, i)" definition ladder_shift :: "ladder \ nat \ nat \ ladder" where "ladder_shift L dn dj = map (deriv_shift dn dj) L" lemma splits_at_append_suffix_prevails: assumes "splits_at (a@b) i u N v" assumes "i < length a" shows "\ v'. v = v'@b \ a=u@[N]@v'" proof - have "min (length a) (Suc i) = Suc i" using Suc_leI assms(2) min.absorb2 by blast then show ?thesis by (metis (no_types) append_assoc append_eq_conv_conj append_take_drop_id assms(1) hd_drop_conv_nth length_take splits_at_def take_hd_drop) qed lemma derivation_shift_right_left_cancel: "derivation_shift (derivation_shift D 0 r) r 0 = D" by (induct D, auto) lemma derivation_shift_left_right_cancel: assumes "derivation_ge D r" shows "derivation_shift (derivation_shift D r 0) 0 r = D" using assms derivation_ge_shift_simp derivation_shift_0_shift by auto lemma LeftDerivation_ge_take: assumes "derivation_ge D k" assumes "LeftDerivation a D b" assumes "D \ []" shows "take k a = take k b \ is_word (take k a)" proof - obtain d D' where d: "d#D' = D" using assms(3) list.exhaust by blast then have "\ x. LeftDerives1 a (fst d) (snd d) x \ LeftDerivation x D' b" using LeftDerivation.simps(2) assms(2) by blast then obtain x where x: "LeftDerives1 a (fst d) (snd d) x \ LeftDerivation x D' b" by blast have fst_d_k: "fst d \ k" using d assms(1) derivation_ge_cons by blast from x fst_d_k have is_word: "is_word (take k a)" by (metis LeftDerives1_def append_take_drop_id is_word_append leftmost_def min.absorb2 take_append take_take) have is_eq: "take k a = take k b" using Derivation_take LeftDerivation_implies_Derivation assms(1) assms(2) by blast show ?thesis using is_word is_eq by blast qed lemma LeftDerivationFix_splits_at_symbol: assumes "LeftDerivationFix a i D j b" shows "\ U a1 a2 b1 b2 n. splits_at a i a1 U a2 \ splits_at b j b1 U b2 \ n \ length D \ LeftDerivation a1 (take n D) b1 \ derivation_ge (drop n D) (Suc(length b1)) \ LeftDerivation a2 (derivation_shift (drop n D) (Suc(length b1)) 0) b2 \ (n = length D \ (n < length D \ is_word (b1@[U])))" proof - note hyp = LeftDerivationFix_def[where \=a and \=b and D=D and i=i and j=j] from hyp obtain E F where EF: "D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take i a) E (take j b) \ LeftDerivation (drop (Suc i) a) F (drop (Suc j) b)" using assms by blast have "\ \. LeftDerivation a E \ \ LeftDerivation \ (derivation_shift F 0 (Suc j)) b" using EF LeftDerivation_append assms(1) hyp by blast then obtain \ where \_intro: "LeftDerivation a E \ \ LeftDerivation \ (derivation_shift F 0 (Suc j)) b" by blast have "LeftDerivation ((take i a)@(drop i a)) E ((take j b)@(drop i a))" by (metis EF LeftDerivation_append_suffix append_take_drop_id assms(1) hyp is_sentence_concat) then have "LeftDerivation a E ((take j b)@(drop i a))" by simp then have \_decomposed: "\ = (take j b)@(drop i a)" using Derivation_unique_dest LeftDerivation_implies_Derivation \_intro by blast have derivation: "LeftDerivation \ (derivation_shift F 0 (Suc j)) b" by (simp add: \_intro) have "\ n. n \ length D \ E = take n D" by (metis EF append_eq_conv_conj is_prefix_length is_prefix_take) then obtain n where n: "n \ length D \ E = take n D" by blast have F_def: "drop n D = derivation_shift F 0 (Suc j)" by (metis EF append_eq_conv_conj length_take min.absorb2 n) have min_j: "min (length b) j = j" using assms hyp by linarith have derivation_ge_Suc_j: "derivation_ge (drop n D) (Suc j)" using F_def derivation_ge_shift by simp have LeftDerivation_\_b: "LeftDerivation \ (drop n D) b" by (simp add: F_def \_intro) have is_word_Suc_j_b: "n \ length D \ is_word (take (Suc j) b)" by (metis EF F_def LeftDerivation_ge_take \_intro append_Nil2 derivation_ge_Suc_j length_take min.absorb2 n) have take_Suc_j_b_decompose: "take (Suc j) b = take j b @ [a ! i]" using assms hyp take_Suc_conv_app_nth by auto show ?thesis apply (rule_tac x="a ! i" in exI) apply (rule_tac x="take i a" in exI) apply (rule_tac x="drop (Suc i) a" in exI) apply (rule_tac x="take j b" in exI) apply (rule_tac x="drop (Suc j) b" in exI) apply (rule_tac x="n" in exI) apply (auto simp add: min_j) using assms hyp splits_at_def apply blast using assms hyp splits_at_def apply blast using n apply blast using EF n apply simp using F_def apply simp using derivation_ge_shift apply blast using F_def derivation_shift_right_left_cancel apply simp using EF apply blast using n apply arith using is_word_Suc_j_b take_Suc_j_b_decompose is_word_append apply simp+ done qed lemma LeftDerivation_breakdown': "LeftDerivation (u @ v) D w \ \n w1 w2. n \ length D \ w = w1 @ w2 \ LeftDerivation u (take n D) w1 \ derivation_ge (drop n D) (length w1) \ LeftDerivation v (derivation_shift (drop n D) (length w1) 0) w2" proof - assume hyp: "LeftDerivation (u @ v) D w" from LeftDerivation_breakdown[OF hyp] obtain n w1 w2 where breakdown: "w = w1 @ w2 \ LeftDerivation u (take n D) w1 \ derivation_ge (drop n D) (length w1) \ LeftDerivation v (derivation_shift (drop n D) (length w1) 0) w2" by blast obtain m where m: "m = min (length D) n" by blast have take_m: "take m D = take n D" using m is_prefix_take take_prefix by fastforce have drop_m: "drop m D = drop n D" by (metis append_eq_conv_conj append_take_drop_id length_take m) have m_bound: "m \ length D" by (simp add: m) show ?thesis apply (rule_tac x="m" in exI) apply (rule_tac x="w1" in exI) apply (rule_tac x="w2" in exI) using breakdown m_bound take_m drop_m by auto qed lemma LeftDerives1_append_replace_in_left: assumes ld1: "LeftDerives1 (\@\) i r \" assumes i_bound: "i < length \" shows "\ \'. \ = \'@\ \ LeftDerives1 \ i r \' \ i + length (snd r) \ length \'" proof - obtain \' where \': "\' = (take i \)@(snd r)@(drop (i+1) \)" by blast have fst_r: "fst r = \ ! i" proof - have "\ss n p ssa. \ LeftDerives1 ss n p ssa \ Derives1 ss n p ssa" using LeftDerives1_implies_Derives1 by blast then have "Derives1 (\ @ \) i r \" using ld1 by blast then show ?thesis using Derives1_nonterminal i_bound splits_at_def by auto qed have "Derives1 \ i r \'" using i_bound ld1 apply (auto simp add: \' Derives1_def) apply (rule_tac x="take i \" in exI) apply (rule_tac x="drop (i+1) \" in exI) apply (rule_tac x="fst r" in exI) apply auto apply (simp add: fst_r) using id_take_nth_drop apply blast using Derives1_sentence1 LeftDerives1_implies_Derives1 is_sentence_concat is_sentence_take apply blast apply (metis Derives1_sentence1 LeftDerives1_implies_Derives1 append_take_drop_id is_sentence_concat) using Derives1_rule LeftDerives1_implies_Derives1 by blast then have leftderives1_\_\': "LeftDerives1 \ i r \'" using LeftDerives1_def i_bound ld1 leftmost_cons_less by auto have i_bound_\': "i + length (snd r) \ length \'" using \' i_bound by (simp add: add_mono_thms_linordered_semiring(2) le_add1 less_or_eq_imp_le min.absorb2) have is_sentence_\: "is_sentence \" using Derives1_sentence1 LeftDerives1_implies_Derives1 is_sentence_concat ld1 by blast then have \: "\ = \'@\" using ld1 leftderives1_\_\' Derives1_append_suffix Derives1_unique_dest LeftDerives1_implies_Derives1 by blast show ?thesis apply (rule_tac x="\'" in exI) using \ i_bound_\' leftderives1_\_\' by blast qed lemma LeftDerivationIntro_propagate: assumes intro: "LeftDerivationIntro (\@\) i r ix D j \" assumes i_\: "i < length \" assumes non: "is_nonterminal (\ ! j)" shows "\ \. LeftDerivation \ ((i,r)#D) \ \ \ = \@\ \ j < length \" proof - from intro LeftDerivationIntro_def[where \="\@\" and i=i and r=r and ix=ix and D=D and j=j and \=\] obtain \ where ld_\: "LeftDerives1 (\ @ \) i r \" and ix: "ix < length (snd r) \ snd r ! ix = \ ! j" and \_fix: "LeftDerivationFix \ (i + ix) D j \" by blast from LeftDerives1_append_replace_in_left[OF ld_\ i_\] obtain \' where \': "\ = \' @ \ \ LeftDerives1 \ i r \' \ i + length (snd r) \ length \'" by blast have i_plus_ix_bound: "i + ix < length \'" using \' ix by linarith have ld_\: "LeftDerivationFix (\' @ \) (i + ix) D j \" using \_fix \' by simp then have non_i_ix: "is_nonterminal ((\' @ \) ! (i + ix))" by (simp add: LeftDerivationFix_def non) from LeftDerivationFix_splits_at_nonterminal[OF ld_\ non_i_ix] obtain U a1 a2 b1 where U: "splits_at (\' @ \) (i + ix) a1 U a2 \ splits_at \ j b1 U a2 \ LeftDerivation a1 D b1" by blast have "\ q. a2 = q@\ \ \' = a1 @ [U] @ q" using splits_at_append_suffix_prevails[OF _ i_plus_ix_bound, where b=\] U by blast then obtain q where q: "a2 = q@\ \ \' = a1 @ [U] @ q" by blast show ?thesis apply (rule_tac x="b1@[U]@q" in exI) apply auto apply (rule_tac x="\'" in exI) apply (metis LeftDerivationFix_def LeftDerivation_append_suffix U \' q append_Cons append_Nil is_sentence_concat ld_\) using U q splits_at_combine apply auto[1] using U splits_at_def by auto qed lemma LeftDerivationIntro_finish: assumes intro: "LeftDerivationIntro (\@\) i r ix D j \" assumes i_\: "i < length \" shows "\ k \ \'. k \ length D \ LeftDerivation \ ((i, r)#(take k D)) \ \ LeftDerivation (\ @ \) ((i, r)#(take k D)) (\ @ \) \ derivation_ge (drop k D) (length \) \ LeftDerivation \ (derivation_shift (drop k D) (length \) 0) \' \ \ = \ @ \' \ j < length \" proof - from intro LeftDerivationIntro_def[where \="\@\" and i=i and r=r and ix=ix and D=D and j=j and \=\] obtain \ where ld_\: "LeftDerives1 (\ @ \) i r \" and ix: "ix < length (snd r) \ snd r ! ix = \ ! j" and \_fix: "LeftDerivationFix \ (i + ix) D j \" by blast from LeftDerives1_append_replace_in_left[OF ld_\ i_\] obtain \' where \': "\ = \' @ \ \ LeftDerives1 \ i r \' \ i + length (snd r) \ length \'" by blast have i_plus_ix_bound: "i + ix < length \'" using \' ix by linarith have ld_\: "LeftDerivationFix (\' @ \) (i + ix) D j \" using \_fix \' by simp from LeftDerivationFix_splits_at_symbol[OF ld_\] obtain U a1 a2 b1 b2 n where U: "splits_at (\' @ \) (i + ix) a1 U a2 \ splits_at \ j b1 U b2 \ n \ length D \ LeftDerivation a1 (take n D) b1 \ derivation_ge (drop n D) (Suc (length b1)) \ LeftDerivation a2 (derivation_shift (drop n D) (Suc (length b1)) 0) b2 \ (n = length D \ n < length D \ is_word (b1 @ [U]))" by blast have n_bound: "n \ length D" using U by blast have "\ q. a2 = q@\ \ \' = a1 @ [U] @ q" using splits_at_append_suffix_prevails[OF _ i_plus_ix_bound, where b=\] U by blast then obtain q where q: "a2 = q@\ \ \' = a1 @ [U] @ q" by blast have j: "j = length b1" using U by (simp add: dual_order.strict_implies_order min.absorb2 splits_at_def) have "n = length D \ n < length D \ is_word (b1 @ [U])" using U by blast then show ?thesis proof (induct rule: disjCases2) case 1 from 1 have drop_n_D: "drop n D = []" by (simp add: U) then have "LeftDerivation a2 [] b2" using U by simp then have a2_eq_b2: "a2 = b2" by simp show ?case apply (rule_tac x="n" in exI) apply (rule_tac x="b1@[U]@q" in exI) apply (rule_tac x="\" in exI) apply auto apply (simp add: 1) apply (rule_tac x="\'" in exI) apply (metis LeftDerivationFix_is_sentence LeftDerivation_append_suffix U \' append_Cons append_Nil is_sentence_concat ld_\ q) apply (rule_tac x="\' @ \" in exI) apply (metis "1.hyps" LeftDerivationFix_def U \' a2_eq_b2 id_take_nth_drop ld_\ ld_\ q splits_at_def take_all) apply (simp add: drop_n_D)+ apply (metis U a2_eq_b2 id_take_nth_drop q splits_at_def) using j apply arith done next case 2 obtain E where E: "E = (derivation_shift (drop n D) (Suc (length b1)) 0)" by blast then have "LeftDerivation (q@\) E b2" using U q by simp from LeftDerivation_breakdown'[OF this] obtain n' w1 w2 where w1w2: "n' \ length E \ b2 = w1 @ w2 \ LeftDerivation q (take n' E) w1 \ derivation_ge (drop n' E) (length w1) \ LeftDerivation \ (derivation_shift (drop n' E) (length w1) 0) w2" by blast have length_E_D: "length E = length D - n" using E n_bound by simp have n_plus_n'_bound: "n + n' \ length D" using length_E_D w1w2 n_bound by arith have take_breakdown: "take (n + n') D = (take n D) @ (take n' (drop n D))" using take_add by blast have q_w1: "LeftDerivation q (take n' E) w1" using w1w2 by blast have isw: "is_word (b1 @ [U])" using 2 by blast have take_n': "take n' (drop n D) = derivation_shift (take n' E) 0 (Suc (length b1))" by (metis E U derivation_shift_left_right_cancel take_derivation_shift) have \'_derives_b1_U_w1: "LeftDerivation \' (take (n + n') D) (b1 @ U # w1)" apply (subst take_breakdown) apply (rule_tac LeftDerivation_implies_append[where b="b1@[U]@q"]) apply (metis LeftDerivationFix_is_sentence LeftDerivation_append_suffix U is_sentence_concat ld_\ q) apply (simp add: take_n') by (rule LeftDerivation_append_prefix[OF q_w1, where u = "b1@[U]", OF isw, simplified]) have dge: "derivation_ge (drop (n + n') D) (Suc (length b1 + length w1))" proof - have "derivation_ge (drop n' (drop n D)) (length b1 + 1 + length w1)" by (metis (no_types) E Suc_eq_plus1 U append_take_drop_id derivation_ge_append derivation_ge_shift_plus drop_derivation_shift w1w2) then show "derivation_ge (drop (n + n') D) (Suc (length b1 + length w1))" by (metis (no_types) Suc_eq_plus1 add.commute drop_drop semiring_normalization_rules(23)) qed show ?case apply (rule_tac x="n+n'" in exI) apply (rule_tac x="b1 @ [U] @ w1" in exI) apply (rule_tac x=w2 in exI) apply auto using n_plus_n'_bound apply simp apply (rule_tac x="\'" in exI) using \' \'_derives_b1_U_w1 apply blast apply (rule_tac x="\' @ \" in exI) apply (metis Cons_eq_appendI LeftDerivationFix_is_sentence LeftDerivation_append_suffix \' \'_derives_b1_U_w1 append_assoc is_sentence_concat ld_\ ld_\) apply (rule dge) apply (metis E Suc_eq_plus1 add.commute derivation_shift_0_shift drop_derivation_shift drop_drop w1w2) using U splits_at_combine w1w2 apply auto[1] by (simp add: j) qed qed lemma LeftDerivationLadder_propagate: "LeftDerivationLadder (\@\) D L \ \ ladder_i L 0 < length \ \ n = ladder_n L index \ index < length L \ if (index + 1 < length L) then (\ \. LeftDerivation \ (take n D) \ \ ladder_\ (\@\) D L index = \@\ \ ladder_j L index < length \) else (\ n' \ \'. (index = 0 \ ladder_prev_n L index < n') \ n' \ n \ LeftDerivation \ (take n' D) \ \ LeftDerivation (\@\) (take n' D) (\@\) \ derivation_ge (drop n' D) (length \) \ LeftDerivation \ (derivation_shift (drop n' D) (length \) 0) \' \ ladder_\ (\@\) D L index = \@\' \ ladder_j L index < length \)" proof (induct index arbitrary: n) case 0 have ldfix: "LeftDerivationFix (\@\) (ladder_i L 0) (take n D) (ladder_j L 0) (ladder_\ (\@\) D L 0)" using "0.prems"(1) "0.prems"(3) LeftDerivationLadder_def by blast from 0 have "1 < length L \ 1 = length L" by arith then show ?case proof (induct rule: disjCases2) case 1 have "LeftDerivationIntrosAt (\@\) D L 1" using "0.prems"(1) "1.hyps" LeftDerivationIntros_def LeftDerivationLadder_def by blast from LeftDerivationIntrosAt_implies_nonterminal[OF this] have "is_nonterminal (ladder_\ (\ @ \) D L 0 ! ladder_j L 0)" by (simp add: ladder_\_def ladder_i_def) with ldfix have "is_nonterminal ((\@\) ! (ladder_i L 0))" by (simp add: LeftDerivationFix_def) from LeftDerivationFix_splits_at_nonterminal[OF ldfix this] obtain U a1 a2 b where thesplit: "splits_at (\ @ \) (ladder_i L 0) a1 U a2 \ splits_at (ladder_\ (\ @ \) D L 0) (ladder_j L 0) b U a2 \ LeftDerivation a1 (take n D) b" by blast have "\ \'. a2 = \' @ \ \ \ = a1 @ [U] @ \'" using thesplit splits_at_append_suffix_prevails using "0.prems"(2) by blast then obtain \' where \': "a2 = \' @ \ \ \ = a1 @ ([U] @ \')" by blast obtain \ where \: "\ = b @ ([U] @ \')" by blast have "is_sentence \" using LeftDerivationFix_is_sentence is_sentence_concat ldfix by blast then have "is_sentence ([U] @ \')" using \' is_sentence_concat by blast with \' thesplit have "LeftDerivation (a1 @ ([U] @ \')) (take n D) (b @ ([U] @ \'))" using LeftDerivation_append_suffix by blast then have \_derives_\: "LeftDerivation \ (take n D) \" using \ \' by blast have \_append_\: "ladder_\ (\ @ \) D L 0 = \@\" by (metis \ \' append_assoc splits_at_combine thesplit) have ladder_j_bound: "ladder_j L 0 < length \" by (metis One_nat_def \ diff_add_inverse dual_order.strict_implies_order leD le_add1 length_Cons length_append length_take list.size(3) min.absorb2 neq0_conv splits_at_def thesplit zero_less_diff zero_less_one) show ?case using 1 apply simp apply (rule_tac x="\" in exI) by (auto simp add: \_derives_\ \_append_\ ladder_j_bound) next case 2 note case_2 = 2 have n_def: "n = length D" by (metis "0.prems"(1) "0.prems"(3) "2.hyps" LeftDerivationLadder_def One_nat_def diff_Suc_1 is_ladder_def ladder_last_n_intro) then have take_n_D: "take n D = D" by (simp add: eq_imp_le) from LeftDerivationFix_splits_at_symbol[OF ldfix] obtain U a1 a2 b1 b2 m where U: "splits_at (\ @ \) (ladder_i L 0) a1 U a2 \ splits_at (ladder_\ (\ @ \) D L 0) (ladder_j L 0) b1 U b2 \ m \ length (take n D) \ LeftDerivation a1 (take m (take n D)) b1 \ derivation_ge (drop m (take n D)) (Suc (length b1)) \ LeftDerivation a2 (derivation_shift (drop m (take n D)) (Suc (length b1)) 0) b2 \ (m = length (take n D) \ (m < length (take n D) \ is_word (b1 @ [U])))" by blast obtain D' where D': "D' = derivation_shift (drop m D) (Suc (length b1)) 0" by blast then have a2_derives_b2: "LeftDerivation a2 D' b2" using U take_n_D by auto from U have m_leq_n: "m \ n" by (simp add: "0.prems"(1) "0.prems"(3) "0.prems"(4) LeftDerivationLadder_def is_ladder_def min.absorb2) from U have "splits_at (\ @ \) (ladder_i L 0) a1 U a2" by blast from splits_at_append_suffix_prevails[OF this 0(2)] obtain v' where v': "a2 = v' @ \ \ \ = a1 @ [U] @ v'" by blast have a1_derives_b1: "LeftDerivation a1 (take m D) b1" using m_leq_n U by (metis "0.prems"(1) "0.prems"(3) "2.hyps" LeftDerivationLadder_def One_nat_def cancel_comm_monoid_add_class.diff_cancel is_ladder_def ladder_last_n_intro order_refl take_all) have "LeftDerivation (v' @ \) D' b2" using a2_derives_b2 v' by simp from LeftDerivation_breakdown'[OF this] obtain m' w1 w2 where w12: "b2 = w1 @ w2 \ m' \ length D' \ LeftDerivation v' (take m' D') w1 \ derivation_ge (drop m' D') (length w1) \ LeftDerivation \ (derivation_shift (drop m' D') (length w1) 0) w2" by blast have "length D' \ length D - m" by (simp add: D') then have "m' \ length D - m" using w12 dual_order.trans by blast - then have m_m'_leq_n: "m + m' \ n" using n_def m_leq_n Nat.le_diff_conv2 add.commute + then have m_m'_leq_n: "m + m' \ n" using n_def m_leq_n le_diff_conv2 add.commute by linarith obtain \ where \: "\ = b1 @ ([U] @ w1)" by blast have "is_sentence ([U] @ v')" using LeftDerivationFix_is_sentence is_sentence_concat ldfix v' by blast then have "LeftDerivation (a1 @ ([U] @ v')) (take m D) (b1 @ ([U] @ v'))" using LeftDerivation_append_suffix a1_derives_b1 by blast then have \_derives_pre_\: "LeftDerivation \ (take m D) (b1 @ ([U] @ v'))" using v' by blast have "m = n \ (m < n \ is_word (b1 @ [U]))" using U n_def[symmetric] take_n_D by simp then have pre_\_derives_\: "LeftDerivation (b1 @ ([U] @ v')) (take m' (drop m D)) \" proof (induct rule: disjCases2) case 1 then have "m' = 0" using m_m'_leq_n by arith then show ?case apply (simp add: \) using w12 by simp next case 2 then have is_word_prefix: "is_word (b1 @ [U])" by blast have take_drop_eq: "take m' (drop m D) = derivation_shift (take m' D') 0 (length (b1 @ [U]))" apply (simp add: D' take_derivation_shift) by (metis U derivation_shift_left_right_cancel take_derivation_shift take_n_D) have v'_derives_w1: "LeftDerivation v' (take m' D') w1" by (simp add: w12) with is_word_prefix have "LeftDerivation ((b1 @ [U]) @ v') (derivation_shift (take m' D') 0 (length (b1 @ [U]))) ((b1 @ [U]) @ w1)" using LeftDerivation_append_prefix by blast with take_drop_eq show ?case by (simp add: \) qed have "(take m D) @ (take m' (drop m D)) = (take (m + m') D)" by (simp add: take_add) then have \_derives_\: "LeftDerivation \ (take (m + m') D) \" using LeftDerivation_implies_append \_derives_pre_\ pre_\_derives_\ by fastforce have derivation_ge_drop_m_m': "derivation_ge (drop (m + m') D) (length \)" proof - have f1: "drop m' (drop m D) = drop (m + m') D" by (simp add: add.commute) have "derivation_ge (drop m' (drop m D)) (Suc (length b1))" by (metis (no_types) U append_take_drop_id derivation_ge_append take_n_D) then show "derivation_ge (drop (m + m') D) (length \)" using f1 by (metis (no_types) D' \ append_assoc derivation_ge_shift_plus drop_derivation_shift length_append length_append_singleton w12) qed have \_derives_w2: "LeftDerivation \ (derivation_shift (drop (m + m') D) (length \) 0) w2" proof - have "derivation_shift (drop m' D') (length w1) 0 = derivation_shift (drop (m + m') D) (length \) 0" by (simp add: D' \ add.commute derivation_shift_0_shift drop_derivation_shift) then show "LeftDerivation \ (derivation_shift (drop (m + m') D) (length \) 0) w2" using w12 by presburger qed have ladder_\_def: "ladder_\ (\ @ \) D L 0 = \ @ w2" using U \ splits_at_combine w12 by auto have ladder_j_bound: "ladder_j L 0 < length \" using U \ splits_at_def by auto show ?case using 2 apply simp apply (rule_tac x="m + m'" in exI) apply (auto simp add: m_m'_leq_n) apply (rule_tac x="\" in exI) apply (auto simp add: \_derives_\) using LeftDerivationFix_is_sentence LeftDerivation_append_suffix \_derives_\ is_sentence_concat ldfix apply blast using derivation_ge_drop_m_m' apply blast apply (rule_tac x="w2" in exI) apply auto using \_derives_w2 apply blast using ladder_\_def apply blast using ladder_j_bound apply blast done qed next case (Suc index) have step: "LeftDerivationIntrosAt (\@\) D L (Suc index)" by (metis LeftDerivationIntros_def LeftDerivationLadder_def Suc.prems(1) Suc.prems(4) Suc_eq_plus1_left le_add1) have index_plus_1_bound: "index + 1 < length L" using Suc.prems(4) by linarith then have index_bound: "index < length L" by arith obtain n' where n': "n' = ladder_n L index" by blast from Suc.hyps[OF Suc.prems(1) Suc.prems(2) n' index_bound] index_plus_1_bound have "\ \'. LeftDerivation \ (take n' D) \' \ ladder_\ (\@\) D L index = \'@\ \ ladder_j L index < length \'" by auto then obtain \' where \': "LeftDerivation \ (take n' D) \' \ ladder_\ (\@\) D L index = \'@\ \ ladder_j L index < length \'" by blast have Suc_index_bound: "Suc index < length L" using Suc.prems by auto have is_ladder: "is_ladder D L" using Suc.prems LeftDerivationLadder_def by auto have n_def: "n = ladder_n L (Suc index)" using Suc_index_bound n' by (simp add: Suc.prems(3)) with n' have n'_less_n: "n' < n" using is_ladder Suc_index_bound is_ladder_def lessI by blast have ladder_\_is_\: "ladder_\ (\@\) D L (Suc index) = ladder_\ (\@\) D L index" by (simp add: ladder_\_def) obtain i where i: "i = ladder_i L (Suc index)" by blast obtain e where e: "e = (D ! n')" by blast obtain E where E: "E = drop (Suc n') (take n D)" by blast obtain ix where ix: "ix = ladder_ix L (Suc index)" by blast obtain j where j: "j = ladder_j L (Suc index)" by blast obtain \ where \: "\ = ladder_\ (\@\) D L (Suc index)" by blast have intro: "LeftDerivationIntro (\'@\) i (snd e) ix E j \" by (metis E LeftDerivationIntrosAt_def \' \ ladder_\_is_\ diff_Suc_1 e i ix j local.step n' n_def) have is_eq_fst_e: "i = fst e" by (metis LeftDerivationIntrosAt_def diff_Suc_1 e i local.step n') have i_less_\': "i < length \'" using i \' ladder_i_def by simp have "(Suc index) + 1 < length L \ (Suc index) + 1 = length L" using Suc_index_bound by arith then show ?case proof (induct rule: disjCases2) case 1 from 1 have "LeftDerivationIntrosAt (\@\) D L (Suc (Suc index))" by (metis LeftDerivationIntros_def LeftDerivationLadder_def Suc.prems(1) Suc_eq_plus1 Suc_eq_plus1_left le_add1) from LeftDerivationIntrosAt_implies_nonterminal[OF this] have "is_nonterminal (ladder_\ (\ @ \) D L (Suc (Suc index)) ! ladder_i L (Suc (Suc index)))" by blast then have non_\_j: "is_nonterminal (\ ! j)" by (simp add: \ j ladder_\_def ladder_i_def) from LeftDerivationIntro_propagate[OF intro i_less_\' non_\_j] obtain \ where \: "LeftDerivation \' ((i, snd e) # E) \ \ \ = \ @ \ \ j < length \" by blast have \_\: "LeftDerivation \ ((take n' D)@((i, snd e) # E)) \" using \' \ LeftDerivation_implies_append by blast have i_e: "(i, snd e) = e" by (simp add: is_eq_fst_e) have take_n_D_e: "((take n' D)@(e # E)) = take n D" proof - (* automatically found *) obtain nn :: "(nat \ symbol \ symbol list) list \ (nat \ nat \ nat) list \ nat" and nna :: "(nat \ symbol \ symbol list) list \ (nat \ nat \ nat) list \ nat" and nnb :: "(nat \ symbol \ symbol list) list \ (nat \ nat \ nat) list \ nat" where f1: "(\ps psa. \ is_ladder ps psa \ psa \ [] \ (\n. \ n < length psa \ ladder_n psa n \ length ps) \ (\n na. (\ n < na \ \ na < length psa) \ ladder_n psa n < ladder_n psa na) \ ladder_last_n psa = length ps) \ (\ps psa. (psa = [] \ nn ps psa < length psa \ \ ladder_n psa (nn ps psa) \ length ps \ (nna ps psa < nnb ps psa \ nnb ps psa < length psa) \ \ ladder_n psa (nna ps psa) < ladder_n psa (nnb ps psa) \ ladder_last_n psa \ length ps) \ is_ladder ps psa)" using is_ladder_def by moura then have f2: "ladder_last_n L = length D" using is_ladder by blast have f3: "min (ladder_last_n L) n = n" using f1 by (metis (no_types) Suc_eq_plus1 index_plus_1_bound is_ladder min.absorb2 n_def) then have "take n' (take n D) @ take n D ! n' # E = take n D" using f2 by (metis E id_take_nth_drop length_take n'_less_n) then show ?thesis using f3 f2 by (metis (no_types) append_assoc append_eq_conv_conj dual_order.strict_implies_order e length_take min.absorb2 n'_less_n nth_append) qed from 1 show ?case apply auto apply (rule_tac x=\ in exI) apply auto using \_\ i_e take_n_D_e apply auto[1] using \ \ apply blast using \ j by blast next case 2 from LeftDerivationIntro_finish[OF intro i_less_\'] obtain k \ \' where kw\': "k \ length E \ LeftDerivation \' ((i, snd e) # take k E) \ \ LeftDerivation (\' @ \) ((i, snd e) # take k E) (\ @ \) \ derivation_ge (drop k E) (length \) \ LeftDerivation \ (derivation_shift (drop k E) (length \) 0) \' \ \ = \ @ \' \ j < length \" by blast have ladder_last_n_1: "ladder_last_n L = n" by (metis "2.hyps" Suc_eq_plus1 diff_Suc_1 ladder_last_n_def n_def) from is_ladder have ladder_last_n_2: "ladder_last_n L = length D" using is_ladder_def by blast from ladder_last_n_1 ladder_last_n_2 have n_eq_length_D: "n = length D" by blast have take_split: "take (Suc (n' + k)) D = (take n' D) @ ((i, snd e) # take k E)" apply (simp add: E n_eq_length_D) by (metis (no_types, lifting) Cons_eq_appendI add_Suc append_eq_appendI e is_eq_fst_e n'_less_n n_eq_length_D prod.collapse self_append_conv2 take_Suc_conv_app_nth take_add) have \_\: "LeftDerivation \ (take (Suc (n' + k)) D) \" apply (subst take_split) apply (rule LeftDerivation_implies_append[where b="\'"]) apply (simp add: \') using kw\' by blast have Suc_n'_k_bound: "Suc (n' + k) \ n" using E kw\' n'_less_n by auto[1] from 2 show ?case apply auto apply (rule_tac x="Suc (n' + k)" in exI) apply auto apply (simp add: ladder_prev_n_def n') using Suc_n'_k_bound apply blast apply (rule_tac x="\" in exI) apply auto using \_\ apply blast using \_\ LeftDerivationFix_def LeftDerivationLadder_def LeftDerivation_append_suffix Suc.prems(1) is_sentence_concat apply auto[1] apply (metis E add.commute add_Suc_right drop_drop kw\' n_eq_length_D nat_le_linear take_all) apply (rule_tac x="\'" in exI) apply auto apply (metis E LeftDerivationLadder_ladder_n_bound Suc.prems(1) Suc_index_bound add.commute add_Suc_right drop_drop kw\' n_def n_eq_length_D take_all) using \ kw\' apply blast using j kw\' by blast qed qed lemma ladder_i_of_cut_at_0: assumes L_non_empty: "L \ []" shows "ladder_i (ladder_cut L n) 0 = ladder_i L 0" proof - have "length L \ 0" using L_non_empty by auto then have "length L = 1 \ length L > 1" by arith then show ?thesis proof (induct rule: disjCases2) case 1 then show ?case apply (simp add: ladder_cut_def ladder_i_def deriv_i_def) by (simp add: assms hd_conv_nth) next case 2 then show ?case apply (simp add: ladder_cut_def ladder_i_def deriv_i_def) by (metis diff_is_0_eq hd_conv_nth leD list_update_nonempty nth_list_update_neq) qed qed lemma ladder_last_j_of_cut: assumes L_non_empty: "L \ []" shows "ladder_last_j (ladder_cut L n) = ladder_last_j L" proof - have length_L_nonzero: "length L \ 0" using L_non_empty by auto then have length_ladder_cut: "length (ladder_cut L n) = length L" by (metis ladder_cut_def length_list_update) show ?thesis apply (simp add: ladder_last_j_def length_ladder_cut) apply (simp add: ladder_cut_def ladder_j_def deriv_j_def) by (metis length_L_nonzero diff_less neq0_conv nth_list_update_eq snd_conv zero_less_Suc) qed lemma length_ladder_cut: assumes L_non_empty: "L \ []" shows "length (ladder_cut L n) = length L" by (metis ladder_cut_def length_list_update) lemma ladder_last_n_of_cut: assumes L_non_empty: "L \ []" shows "ladder_last_n (ladder_cut L n) = n" proof - show ?thesis apply (simp add: ladder_last_n_def length_ladder_cut[OF L_non_empty]) apply (simp add: ladder_n_def ladder_cut_def deriv_n_def) by (metis assms diff_Suc_less fst_conv length_greater_0_conv nth_list_update_eq) qed lemma ladder_n_of_cut: assumes L_non_empty: "L \ []" assumes "index < length L - 1" shows "ladder_n (ladder_cut L n) index = ladder_n L index" by (metis assms(2) ladder_cut_def ladder_n_def nat_neq_iff nth_list_update_neq) lemma ladder_n_prev_bound: assumes ladder: "is_ladder D L" assumes u_bound: "u < length L - 1" shows "ladder_n L u \ ladder_prev_n L (length L - 1)" proof - have "ladder_n L u \ ladder_n L (length L - 2)" proof - have "u < Suc (length L - 2)" using u_bound by linarith then show ?thesis by (metis (no_types) diff_Suc_less is_ladder_def ladder leI length_greater_0_conv not_less_eq numeral_2_eq_2 order.order_iff_strict) qed then show ?thesis by (metis One_nat_def Suc_diff_Suc diff_Suc_1 ladder_prev_n_def neq0_conv not_less0 numeral_2_eq_2 u_bound zero_less_diff) qed lemma ladder_n_last_is_length: assumes "is_ladder D L" shows "ladder_n L (length L - 1) = length D" using assms is_ladder_def ladder_last_n_intro by auto lemma derivation_ge_shift_implies_derivation_ge: assumes dge: "derivation_ge (derivation_shift F 0 j) k" shows "derivation_ge F (k - j)" proof - have "\ i r. (i, r) \ set (derivation_shift F 0 j) \ i \ k" using dge derivation_ge_def by auto { fix i :: nat fix r :: "symbol \ (symbol list)" assume ir: "(i, r) \ set F" then have "(i + j, r) \ set (derivation_shift F 0 j)" proof - have "(i + j, r) \ (\p. (fst p - 0 + j, snd p)) ` set F" by (metis (lifting) ir diff_zero image_eqI prod.collapse prod.inject) then show ?thesis by (simp add: derivation_shift_def) qed then have "i + j \ k" using dge derivation_ge_def by auto then have "i \ k - j" by auto } then show ?thesis using derivation_ge_def by auto qed lemma Derives1_bound': "Derives1 a i r b \ i \ length b" by (metis Derives1_bound Derives1_take append_Nil2 append_take_drop_id drop_eq_Nil dual_order.strict_implies_order length_take min.absorb2 nat_le_linear) lemma LeftDerivation_Derives1_last: assumes "LeftDerivation a D b" assumes "D \ []" shows "Derives1 (Derive a (take (length D - 1) D)) (fst (last D)) (snd (last D)) b" by (metis Derive LeftDerivation_Derive_take_LeftDerives1 LeftDerivation_implies_Derivation LeftDerives1_implies_Derives1 assms(1) assms(2) last_conv_nth le_refl length_0_conv take_all) lemma last_of_prefix_in_set: assumes "n < length E" assumes "D = E@F" shows "last E \ set (drop n D)" proof - have f1: "last (drop n E) = last E" by (simp add: assms(1)) have "drop n E \ []" by (metis (no_types) Cons_nth_drop_Suc assms(1) list.simps(3)) then show ?thesis using f1 by (metis (no_types) append.simps(2) append_butlast_last_id append_eq_conv_conj assms(2) drop_append in_set_dropD insertCI list.set(2)) qed lemma LeftDerivationFix_cut_appendix: assumes ldfix: "LeftDerivationFix (\@\) i D j (\@\')" assumes \_\: "LeftDerivation \ (take n D) \" assumes n_bound: "n \ length D" assumes dge: "derivation_ge (drop n D) (length \)" assumes i_in: "i < length \" assumes j_in: "j < length \" shows "LeftDerivationFix \ i (take n D) j \" proof - from LeftDerivationFix_def[where \="\@\" and i=i and D=D and j=j and \="\@\'"] obtain E F where EF: "is_sentence (\ @ \) \ is_sentence (\ @ \') \ LeftDerivation (\ @ \) D (\ @ \') \ i < length (\ @ \) \ j < length (\ @ \') \ (\ @ \) ! i = (\ @ \') ! j \ D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take i (\ @ \)) E (take j (\ @ \')) \ LeftDerivation (drop (Suc i) (\ @ \)) F (drop (Suc j) (\ @ \'))" using ldfix by auto with i_in j_in have take_i_E_take_j: "LeftDerivation (take i \) E (take j \)" by (simp add: less_or_eq_imp_le) obtain m where m: "m = length E" by blast { assume n_less_m: "n < m" then have E_nonempty: "E \ []" using gr_implies_not0 list.size(3) m by auto have last_E_in_set: "last E \ set (drop n D)" using last_of_prefix_in_set n_less_m m EF by blast obtain k r where kr: "(k, r) = last E" by (metis old.prod.exhaust) have k_lower_bound: "k \ length \" using dge last_E_in_set kr by (metis derivation_ge_def fst_conv) have "\ \'. Derives1 \' k r (take j \)" using LeftDerivation_Derives1_last take_i_E_take_j by (metis E_nonempty kr fst_conv snd_conv) then have "k \ j" by (metis Derives1_bound' j_in length_take less_imp_le_nat min.absorb2) then have k_upper_bound: "k < length \" using j_in by arith from k_lower_bound k_upper_bound have "False" by arith } then have m_le_n: "m \ n" by arith have take_i_E_take_j: "LeftDerivation (take i \) E (take j \)" by (simp add: take_i_E_take_j) have "take n D = E @ (take (n - m) (derivation_shift F 0 (Suc j)))" using EF m m_le_n by auto then have take_n_D: "take n D = E @ (derivation_shift (take (n - m) F) 0 (Suc j))" using take_derivation_shift by auto obtain F' where F': "F' = derivation_shift (take (n - m) F) 0 (Suc j)" by blast have "LeftDerivation ((take i \)@(drop i \)) E ((take j \)@(drop i \))" using take_i_E_take_j by (metis EF LeftDerivation_append_suffix append_take_drop_id is_sentence_concat) then have "LeftDerivation \ E ((take j \)@(drop i \))" by simp with take_n_D have take_j_drop_i: "LeftDerivation ((take j \)@(drop i \)) F' \" using F' by (metis Derivation_unique_dest LeftDerivation_append LeftDerivation_implies_Derivation \_\) have F'_ge: "derivation_ge F' (Suc j)" using F' derivation_ge_shift by blast have drop_append: "drop i \ = [\!i] @ (drop (Suc i) \)" by (simp add: Cons_nth_drop_Suc i_in) have take_append: "take j \ @ [\!i] = take (Suc j) \" by (metis LeftDerivationFix_def i_in j_in ldfix nth_superfluous_append take_Suc_conv_app_nth) have take_drop_Suc: "(take j \)@(drop i \) = (take (Suc j) \)@(drop (Suc i) \)" by (simp add: drop_append take_append) with take_drop_Suc take_j_drop_i have "LeftDerivation ((take (Suc j) \)@(drop (Suc i) \)) F' \" by auto note helper = LeftDerivation_skip_prefix[OF this] have len_take: "length (take (Suc j) \) = Suc j" by (simp add: Suc_leI j_in min.absorb2) have F'_shift: "derivation_shift F' (Suc j) 0 = take (n - m) F" using F' derivation_shift_right_left_cancel by blast have LeftDerivation_drop: "LeftDerivation (drop (Suc i) \) (take (n - m) F) (drop (Suc j) \)" using helper len_take F'_shift F'_ge by auto show ?thesis apply (auto simp add: LeftDerivationFix_def) using LeftDerivationFix_is_sentence is_sentence_concat ldfix apply blast using LeftDerivationFix_is_sentence is_sentence_concat ldfix apply blast using \_\ apply blast using i_in apply blast using j_in apply blast using LeftDerivationFix_def i_in j_in ldfix apply auto[1] apply (rule_tac x=E in exI) apply (rule_tac x="take (n - m) F" in exI) apply auto using take_n_D apply blast using take_i_E_take_j apply blast using LeftDerivation_drop by blast qed lemma LeftDerivationFix_cut_appendix': assumes ldfix: "LeftDerivationFix (\@\) i D j (\@\')" assumes \_\: "LeftDerivation \ D \" assumes i_in: "i < length \" assumes j_in: "j < length \" shows "LeftDerivationFix \ i D j \" proof - obtain n where n: "n = length D" by blast have "LeftDerivationFix \ i (take n D) j \" apply (rule_tac LeftDerivationFix_cut_appendix) apply (rule ldfix) using \_\ n apply auto[1] using n apply auto[1] using n apply auto[1] using i_in apply blast using j_in apply blast done then show ?thesis using n by auto qed lemma LeftDerivationIntro_cut_appendix: assumes ldfix: "LeftDerivationIntro (\@\) i r ix D j (\@\')" assumes \_\: "LeftDerivation \ ((i,r)#(take n D)) \" assumes n_bound: "n \ length D" assumes dge: "derivation_ge (drop n D) (length \)" assumes i_in: "i < length \" assumes j_in: "j < length \" shows "LeftDerivationIntro \ i r ix (take n D) j \" proof - from LeftDerivationIntro_def[where \="\@\" and i=i and r=r and ix=ix and D=D and j=j and \="\@\'"] obtain \ where \: "LeftDerives1 (\ @ \) i r \ \ ix < length (snd r) \ snd r ! ix = (\ @ \') ! j \ LeftDerivationFix \ (i + ix) D j (\ @ \')" using ldfix by blast then have "\ \'. \ = \' @ \ \ i + length (snd r) \ length \'" using i_in using LeftDerives1_append_replace_in_left by blast then obtain \' where \': "\ = \' @ \ \ i + length (snd r) \ length \'" by blast have \_\': "LeftDerives1 \ i r \'" using \' \ using LeftDerives1_skip_suffix i_in by blast from \_\ obtain u where u: "LeftDerives1 \ i r u \ LeftDerivation u (take n D) \" by auto with \_\' have "u = \'" using Derives1_unique_dest LeftDerives1_implies_Derives1 by blast with u have \'_\: "LeftDerivation \' (take n D) \" by auto have ldfix_append: "LeftDerivationFix (\' @ \) (i + ix) D j (\ @ \')" using \' \ by blast have i_plus_ix_bound: "i + ix < length \'" using \ \' using add_lessD1 le_add_diff_inverse less_asym' linorder_neqE_nat nat_add_left_cancel_less by linarith from LeftDerivationFix_cut_appendix[OF ldfix_append \'_\ n_bound dge i_plus_ix_bound j_in] have ldfix: "LeftDerivationFix \' (i + ix) (take n D) j \" . show ?thesis apply (simp add: LeftDerivationIntro_def) apply (rule_tac x="\'" in exI) apply auto using \_\' apply blast using \ apply blast apply (simp add: \ j_in) using ldfix by blast qed lemma LeftDerivationIntro_cut_appendix': assumes ldfix: "LeftDerivationIntro (\@\) i r ix D j (\@\')" assumes \_\: "LeftDerivation \ ((i,r)#D) \" assumes i_in: "i < length \" assumes j_in: "j < length \" shows "LeftDerivationIntro \ i r ix D j \" proof - obtain n where n: "n = length D" by blast have "LeftDerivationIntro \ i r ix (take n D) j \" apply (rule_tac LeftDerivationIntro_cut_appendix) apply (rule ldfix) using \_\ n apply auto[1] using n apply auto[1] using n apply auto[1] using i_in apply blast using j_in apply blast done then show ?thesis using n by auto qed lemma ladder_n_monotone: "is_ladder D L \ u \ v \ v < length L \ ladder_n L u \ ladder_n L v" by (metis is_ladder_def le_neq_implies_less linear not_less) lemma ladder_i_cut: assumes index_bound: "index < length L" shows "ladder_i (ladder_cut L n) index = ladder_i L index" proof - have "index = 0 \ index > 0" by arith then show ?thesis proof (induct rule: disjCases2) case 1 with index_bound have "L \ []" by (simp add: less_numeral_extra(3)) then show ?case using 1 by (simp add: ladder_i_of_cut_at_0) next case 2 then show ?case apply (auto simp add: ladder_i_def ladder_cut_def ladder_j_def deriv_j_def Let_def) using index_bound by auto qed qed lemma ladder_j_cut: assumes index_bound: "index < length L" shows "ladder_j (ladder_cut L n) index = ladder_j L index" by (metis gr_implies_not0 index_bound ladder_cut_def ladder_j_def ladder_last_j_def ladder_last_j_of_cut length_ladder_cut list.size(3) nth_list_update_neq) lemma ladder_ix_cut: assumes index_lower_bound: "index > 0" assumes index_upper_bound: "index < length L" shows "ladder_ix (ladder_cut L n) index = ladder_ix L index" proof - show ?thesis using index_lower_bound apply (simp add: ladder_ix_def deriv_ix_def) by (metis index_upper_bound ladder_cut_def nth_list_update_eq nth_list_update_neq snd_conv) qed lemma LeftDerivation_from_in_between: assumes \_\: "LeftDerivation \ (take u D) \" assumes \_\: "LeftDerivation \ (take v D) \" assumes u_le_v: "u \ v" shows "LeftDerivation \ (drop u (take v D)) \" proof - have take_split: "take v D = take u D @ (drop u (take v D))" by (metis u_le_v add_diff_cancel_left' drop_take le_Suc_ex take_add) then show ?thesis using \_\ \_\ by (metis (no_types, lifting) Derivation_unique_dest LeftDerivation_append LeftDerivation_implies_Derivation) qed lemma LeftDerivationLadder_cut_appendix_helper: assumes LDLadder: "LeftDerivationLadder (\@\) D L \" assumes ladder_i_in_\: "ladder_i L 0 < length \" shows "\ E F \1 \2 L'. D = E@F \ \ = \1 @ \2 \ LeftDerivationLadder \ E L' \1 \ derivation_ge F (length \1) \ LeftDerivation \ (derivation_shift F (length \1) 0) \2 \ L' = ladder_cut L (length E)" proof - have is_ladder_D_L: "is_ladder D L" using LDLadder LeftDerivationLadder_def by blast obtain n where n: "n = ladder_last_n L" by blast then have n_eq_ladder_n: "n = ladder_n L (length L - 1)" using ladder_last_n_def by blast have length_L_nonzero: "length L > 0" using LeftDerivationLadder_def assms(1) is_ladder_def by blast from LeftDerivationLadder_propagate[OF LDLadder ladder_i_in_\ n_eq_ladder_n] obtain n' \ \' where finish: "(length L - 1 = 0 \ ladder_prev_n L (length L - 1) < n') \ n' \ n \ LeftDerivation \ (take n' D) \ \ LeftDerivation (\ @ \) (take n' D) (\ @ \) \ derivation_ge (drop n' D) (length \) \ LeftDerivation \ (derivation_shift (drop n' D) (length \) 0) \' \ ladder_\ (\ @ \) D L (length L - 1) = \ @ \' \ ladder_j L (length L - 1) < length \" using length_L_nonzero by auto obtain E where E: "E = take n' D" by blast obtain F where F: "F = drop n' D" by blast obtain L' where L': "L' = ladder_cut L (length E)" by blast have \_ladder: "\ = ladder_\ (\ @ \) D L (length L - 1)" by (metis Derive LDLadder LeftDerivationLadder_def LeftDerivation_implies_Derivation append_Nil2 append_take_drop_id drop_eq_Nil is_ladder_def ladder_\_def le_refl n n_eq_ladder_n) then have \: "\ = \ @ \'" using finish by auto have "is_sentence (\@\)" using LDLadder LeftDerivationFix_is_sentence LeftDerivationLadder_def by blast then have is_sentence_\: "is_sentence \" using is_sentence_concat by blast have "is_sentence \" using Derivation_implies_derives LDLadder LeftDerivationFix_is_sentence LeftDerivationLadder_def LeftDerivation_implies_Derivation derives_is_sentence by blast then have is_sentence_\: "is_sentence \" using \ is_sentence_concat by blast have length_L': "length L' = length L" by (metis L' ladder_cut_def length_list_update) have ladder_last_n_L': "ladder_last_n L' = length E" using L' ladder_last_n_of_cut length_L_nonzero by blast have length_D_eq_n: "length D = n" using LDLadder LeftDerivationLadder_def is_ladder_def n by auto then have length_E_eq_n': "length E = n'" by (simp add: E finish min.absorb2) { fix u :: nat assume "u < length L'" then have "u < length L' - 1 \ u = length L' - 1" by arith then have "ladder_n L' u \ length E" proof (induct rule: disjCases2) case 1 have u_bound: "u < length L - 1" using 1 by (simp add: length_L') then have L'_eq_L: "ladder_n L' u = ladder_n L u" using L' ladder_n_of_cut length_L_nonzero length_greater_0_conv by blast from u_bound have "ladder_n L u \ ladder_prev_n L (length L - 1)" using ladder_n_prev_bound LeftDerivationLadder_def assms(1) by blast then show ?case using L'_eq_L finish length_E_eq_n' u_bound by linarith next case 2 then have "ladder_n L' u = length E" using ladder_last_n_L' ladder_last_n_def by auto then show ?case by auto qed } note ladder_n_bound = this { fix u :: nat fix v :: nat assume u_less_v: "u < v" assume v_bound: "v < length L'" have "v < length L' - 1 \ v = length L' - 1" using v_bound by arith then have "ladder_n L' u < ladder_n L' v" proof (induct rule: disjCases2) case 1 show ?case using "1.hyps" L' LeftDerivationLadder_def assms(1) is_ladder_def ladder_n_of_cut length_L' u_less_v by auto next case 2 note v_def = 2 have "v = 0 \ v \ 0" by arith then show ?case proof (induct rule: disjCases2) case 1 then show ?case using u_less_v by auto next case 2 then have "ladder_prev_n L (length L - 1) < n'" using finish v_def length_L' by linarith then show ?case by (metis (no_types, lifting) L' LeftDerivationLadder_def assms(1) ladder_last_n_L' ladder_last_n_def ladder_n_of_cut ladder_n_prev_bound le_neq_implies_less length_E_eq_n' length_L' length_greater_0_conv less_trans u_less_v v_def) qed qed } note ladder_n_pairwise_bound = this have is_ladder_E_L': "is_ladder E L'" apply (auto simp add: is_ladder_def ladder_last_n_L') using length_L_nonzero length_L' apply simp using ladder_n_bound apply blast using ladder_n_pairwise_bound by blast { fix index :: nat assume index_bound: "index + 1 < length L" then have index_le: "index < length L" by arith from index_bound have len_L_minus_1: "length L - 1 \ 0" by arith obtain m where m: "m = ladder_n L index" by blast from LeftDerivationLadder_propagate[OF LDLadder ladder_i_in_\ m index_le] obtain \ where \: "LeftDerivation \ (take m D) \ \ ladder_\ (\ @ \) D L index = \ @ \ \ ladder_j L index < length \" using index_bound by auto have L'_Derive: "ladder_\ \ E L' index = Derive \ (take (ladder_n L' index) E)" by (simp add: ladder_\_def) have ladder_n_L'_eq_L: "ladder_n L' index = ladder_n L index" using L' index_bound ladder_n_of_cut length_L_nonzero by auto have "ladder_prev_n L (length L - 1) < n'" using finish len_L_minus_1 by blast then have n'_is_upper_bound: "ladder_n L (length L - 2) < n'" using index_bound by (metis diff_diff_left ladder_prev_n_def len_L_minus_1 one_add_one) have index_upper_bound: "index \ length L - 2" using index_bound by linarith have ladder_n_upper_bound: "ladder_n L index \ ladder_n L (length L - 2)" apply (rule_tac ladder_n_monotone) apply (rule_tac is_ladder_D_L) apply (rule index_upper_bound) using length_L_nonzero by linarith with n'_is_upper_bound have m_le_n': "m \ n'" using dual_order.strict_implies_order le_less_trans m by linarith then have "take m E = take m D" by (metis E le_take_same length_E_eq_n' order_refl take_all) then have take_helper: "(take (ladder_n L' index) E) = take m D" by (simp add: ladder_n_L'_eq_L m) then have Derive_eq_\: "Derive \ (take (ladder_n L' index) E) = \" by (simp add: Derive LeftDerivation_implies_Derivation \) then have t1: "ladder_\ (\@\) D L index = (ladder_\ \ E L' index) @ \" by (simp add: L'_Derive \) have \_eq: "\ = ladder_\ \ E L' index" by (simp add: Derive_eq_\ L'_Derive) then have t2: "LeftDerivation \ (take (ladder_n L index) D) (ladder_\ \ E L' index)" using \ m by blast have t3: "(take (ladder_n L' index) E) = take (ladder_n L index) D" by (simp add: m take_helper) have t4: "ladder_j L index < length (ladder_\ \ E L' index)" using \ \_eq by blast have t5: "E ! (ladder_n L' index) = D ! (ladder_n L index)" using E ladder_n_L'_eq_L ladder_n_upper_bound n'_is_upper_bound by auto note t = t1 t2 t3 t4 t5 } note ladder_early_stage = this { fix index :: nat assume index_bound: "index < length L" have i: "ladder_i L' index = ladder_i L index" using L' ladder_i_cut by (simp add: index_bound) have j: "ladder_j L' index = ladder_j L index" using L' ladder_j_cut by (simp add: index_bound) have ix: "index > 0 \ ladder_ix L' index = ladder_ix L index" using L' ladder_ix_cut by (simp add: index_bound) have \: "ladder_\ (\@\) D L index = (ladder_\ \ E L' index) @ \" by (simp add: index_bound ladder_\_def ladder_early_stage(1)) have i_bound: "index > 0 \ ladder_i L' index < length (ladder_\ \ E L' index)" using i index_bound ladder_\_def ladder_early_stage(4) ladder_i_def by auto note ij = i j ix \ i_bound } note ladder_every_stage = this { fix u :: nat fix v :: nat assume u_le_v: "u \ v" assume v_bound: "v < length L" have "ladder_n L u \ ladder_n L v" using is_ladder_D_L ladder_n_monotone u_le_v v_bound by blast } note ladder_L_n_pairwise_le = this { fix index :: nat assume index_lower_bound: "index > 0" assume index_upper_bound: "index + 1 < length L" note derivation = ladder_early_stage(2) have derivation1: "LeftDerivation \ (take (ladder_n L (index - Suc 0)) D) (ladder_\ \ E L' (index - Suc 0))" using derivation[of "index - Suc 0"] index_lower_bound index_upper_bound using One_nat_def Suc_diff_1 Suc_eq_plus1 le_less_trans lessI less_or_eq_imp_le by linarith have derivation2: "LeftDerivation \ (take (ladder_n L index) D) (ladder_\ \ E L' index)" using derivation[of index] index_upper_bound by blast have ladder_\_is_\[symmetric]: "ladder_\ \ E L' (index - Suc 0) = ladder_\ \ E L' index" using index_lower_bound ladder_\_def by auto have ladder_n_le: "ladder_n L (index - Suc 0) \ ladder_n L index" apply (rule_tac ladder_L_n_pairwise_le) apply arith using index_upper_bound by arith from LeftDerivation_from_in_between[OF derivation1 derivation2 ladder_n_le] ladder_\_is_\ have "LeftDerivation (ladder_\ \ E L' index) (drop (ladder_n L' (index - Suc 0)) (take (ladder_n L' index) E)) (ladder_\ \ E L' index)" by (metis L' Suc_leI add_lessD1 index_lower_bound index_upper_bound ladder_early_stage(3) ladder_n_of_cut le_add_diff_inverse2 length_L_nonzero length_greater_0_conv less_diff_conv) } note LeftDerivation_delta_early = this have LeftDerivationFix_\_0: "LeftDerivationFix \ (ladder_i L' 0) (take (ladder_n L' 0) E) (ladder_j L' 0) (ladder_\ \ E L' 0)" proof - have ldfix: "LeftDerivationFix (\@\) (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ (\@\) D L 0)" using LDLadder LeftDerivationLadder_def by blast have ladder_i_L': "ladder_i L' 0 = ladder_i L 0" using L' ladder_i_of_cut_at_0 length_L_nonzero by blast have ladder_j_L': "ladder_j L' 0 = ladder_j L 0" by (metis L' ladder_cut_def ladder_j_def ladder_last_j_def ladder_last_j_of_cut length_L' length_greater_0_conv nth_list_update_neq) have "length L = 1 \ length L > 1" using length_L_nonzero by linarith then show ?thesis proof (induct rule: disjCases2) case 1 from 1 have ladder_n_L'_0: "ladder_n L' 0 = n'" using diff_is_0_eq' ladder_last_n_L' ladder_last_n_def length_E_eq_n' length_L' less_or_eq_imp_le by auto have take_n'_E: "take n' E = E" by (simp add: length_E_eq_n') from ladder_n_L'_0 take_n'_E have take_ladder_n_L': "take (ladder_n L' 0) E = E" by auto have "ladder_n L 0 = length D" by (simp add: "1.hyps" length_D_eq_n n_eq_ladder_n) then have take_ladder_n_L_0: "take (ladder_n L 0) D = D" by simp have ladder_\_\: "ladder_\ \ E L' 0 = \" apply (simp add: ladder_\_def take_ladder_n_L') by (simp add: Derive E LeftDerivation_implies_Derivation finish) have ladder_j_in_\: "ladder_j L 0 < length \" using finish "1.hyps" by auto have ldfix_1: "LeftDerivationFix (\@\) (ladder_i L 0) D (ladder_j L 0) (\@\')" using "1.hyps" \ \_ladder ldfix take_ladder_n_L_0 by auto then have "LeftDerivationFix \ (ladder_i L 0) E (ladder_j L 0) \" by (simp add: E LeftDerivationFix_cut_appendix finish ladder_i_in_\ ladder_j_in_\ length_D_eq_n) then show ?case by (simp add: ladder_i_L' ladder_j_L' take_ladder_n_L' ladder_\_\) next case 2 have h: "0 + 1 < length L" using "2.hyps" by auto note stage = ladder_early_stage[of 0, OF h] have ldfix0: "LeftDerivationFix (\@\) (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) ((ladder_\ \ E L' 0) @ \)" using ladder_i_L' ladder_j_L' ldfix stage(1) stage(3) by auto from LeftDerivationFix_cut_appendix'[OF ldfix0 stage(2) ladder_i_in_\ stage(4)] show ?case by (simp add: ladder_i_L' ladder_j_L' stage(3)) qed qed { fix index :: nat assume index_bounds: "1 \ index \ index + 1 < length L" have introsAt_appendix: "LeftDerivationIntrosAt (\@\) D L index" using LDLadder LeftDerivationIntros_def LeftDerivationLadder_def add_lessD1 index_bounds by blast have index_plus_1_upper_bound: "index + 1 < length L" using index_bounds by arith note early_stage = ladder_early_stage[of index, OF index_plus_1_upper_bound] have ladder_i_L_index_eq_fst: "ladder_i L index = fst (D ! ladder_n L (index - Suc 0))" using introsAt_appendix LeftDerivationIntrosAt_def index_bounds by (metis One_nat_def) have E_at_D_at: "(E ! ladder_n L' (index - Suc 0)) = (D ! ladder_n L (index - Suc 0))" using ladder_early_stage[of "index - Suc 0"] by (metis One_nat_def add_lessD1 index_bounds le_add_diff_inverse2) then have ladder_i_L'_index_eq_fst: "ladder_i L' index = fst (E ! ladder_n L' (index - Suc 0))" using index_bounds ladder_i_L_index_eq_fst ladder_every_stage(1) le_add1 le_less_trans by auto have same_derivation: "(drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) E)) = (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D))" using L' early_stage(3) index_bounds ladder_n_of_cut length_L_nonzero by auto have deriv_split: "(drop (ladder_n L' (index - Suc 0)) (take (ladder_n L' index) E)) = ((ladder_i L' index, snd (E ! ladder_n L' (index - Suc 0))) # drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) E))" by (metis Cons_nth_drop_Suc One_nat_def Suc_le_lessD add_lessD1 diff_Suc_less index_bounds ladder_i_L'_index_eq_fst ladder_n_bound ladder_n_pairwise_bound length_L' length_take min.absorb2 nth_take prod.collapse) have "LeftDerivationIntrosAt \ E L' index" apply (auto simp add: LeftDerivationIntrosAt_def Let_def) using ladder_i_L'_index_eq_fst apply blast apply (rule_tac LeftDerivationIntro_cut_appendix'[where \=\ and \' = \]) apply (metis E_at_D_at LeftDerivationIntrosAt_def One_nat_def Suc_le_lessD add_lessD1 early_stage(1) index_bounds introsAt_appendix ladder_every_stage(2) ladder_every_stage(3) ladder_every_stage(4) ladder_i_L'_index_eq_fst same_derivation) defer 1 using index_bounds ladder_every_stage(5) apply auto[1] using early_stage(4) index_bounds ladder_every_stage(2) apply auto[1] using LeftDerivation_delta_early deriv_split by (metis One_nat_def Suc_le_eq index_bounds) } note LeftDerivationIntrosAt_early = this { fix index :: nat assume index_bounds: "1 \ index \ index + 1 = length L" have introsAt_appendix: "LeftDerivationIntrosAt (\@\) D L index" using LDLadder LeftDerivationIntros_def LeftDerivationLadder_def add_lessD1 index_bounds by (metis Suc_eq_plus1 not_less_eq) have ladder_i_L_index_eq_fst: "ladder_i L index = fst (D ! ladder_n L (index - Suc 0))" using introsAt_appendix LeftDerivationIntrosAt_def index_bounds by (metis One_nat_def) have E_at_D_at: "(E ! ladder_n L' (index - Suc 0)) = (D ! ladder_n L (index - Suc 0))" using ladder_early_stage[of "index - Suc 0"] by (metis One_nat_def Suc_eq_plus1 index_bounds le_add_diff_inverse2 lessI) then have ladder_i_L'_index_eq_fst: "ladder_i L' index = fst (E ! ladder_n L' (index - Suc 0))" using index_bounds ladder_i_L_index_eq_fst ladder_every_stage(1) le_add1 le_less_trans by auto obtain D' where D': "D' = (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D))" by blast obtain k where k: "k = (ladder_n L' index) - (Suc (ladder_n L' (index - Suc 0)))" by blast have ladder_n_L'_index: "ladder_n L' index = length E" by (metis diff_add_inverse2 index_bounds ladder_last_n_L' ladder_last_n_def length_L') have take_is_E: "take (ladder_n L' index) E = E" by (simp add: ladder_n_L'_index) have ladder_n_L_index: "ladder_n L index = length D" by (metis diff_add_inverse2 index_bounds length_D_eq_n n_eq_ladder_n) have take_is_D: "take (ladder_n L index) D = D" by (simp add: ladder_n_L_index) have write_as_take_k_D': "(drop (Suc (ladder_n L' (index - Suc 0))) E) = take k D'" using take_is_D by (metis D' E L' One_nat_def Suc_le_lessD add_diff_cancel_right' diff_Suc_less drop_take index_bounds k ladder_n_L'_index ladder_n_of_cut length_E_eq_n' length_L_nonzero length_greater_0_conv) have k_bound: "k \ length D'" by (metis le_iff_add append_take_drop_id k ladder_n_L'_index length_append length_drop write_as_take_k_D') have D'_alt: "D' = drop (Suc (ladder_n L (index - Suc 0))) D" by (simp add: D' take_is_D) have "LeftDerivationIntrosAt \ E L' index" apply (auto simp add: LeftDerivationIntrosAt_def Let_def) using ladder_i_L'_index_eq_fst apply blast apply (subst take_is_E) apply (subst write_as_take_k_D') apply (rule_tac LeftDerivationIntro_cut_appendix[where \=\ and \' = \']) apply (metis D' Derive E E_at_D_at LeftDerivationIntrosAt_def LeftDerivation_implies_Derivation One_nat_def Suc_le_lessD add_diff_cancel_right' diff_Suc_less finish index_bounds introsAt_appendix ladder_\_def ladder_every_stage(2) ladder_every_stage(3) ladder_every_stage(4) ladder_i_L'_index_eq_fst length_L_nonzero take_is_E) apply (metis Cons_nth_drop_Suc E L' LeftDerivation_from_in_between LeftDerivation_take_derive One_nat_def Suc_le_lessD add_diff_cancel_right' diff_Suc_less finish index_bounds ladder_\_def ladder_\_def ladder_i_L'_index_eq_fst ladder_n_L'_index ladder_n_of_cut ladder_prev_n_def length_E_eq_n' length_L_nonzero less_imp_le_nat less_numeral_extra(3) list.size(3) prod.collapse take_is_E write_as_take_k_D') using k_bound apply blast using D'_alt apply (metis (no_types, lifting) Derive E L' LeftDerivation_implies_Derivation One_nat_def Suc_leI Suc_le_lessD add_diff_cancel_right' diff_Suc_less drop_drop finish index_bounds k ladder_\_def ladder_n_L'_index ladder_n_of_cut ladder_prev_n_def le_add_diff_inverse2 length_E_eq_n' length_L_nonzero length_greater_0_conv less_not_refl2 take_is_E) using index_bounds ladder_every_stage(5) apply auto[1] by (metis Derive E LeftDerivation_implies_Derivation One_nat_def add_diff_cancel_right' diff_Suc_less finish index_bounds ladder_\_def ladder_every_stage(2) length_L_nonzero take_is_E) } note LeftDerivationIntrosAt_last = this have ladder_E_L': "LeftDerivationLadder \ E L' \" apply (auto simp add: LeftDerivationLadder_def) using finish E apply blast using is_ladder_E_L' apply blast using LeftDerivationFix_\_0 apply blast using LeftDerivationIntros_def LeftDerivationIntrosAt_early LeftDerivationIntrosAt_last by (metis Suc_eq_plus1 Suc_leI le_neq_implies_less length_L') show ?thesis apply (rule_tac x=E in exI) apply (rule_tac x=F in exI) apply (rule_tac x=\ in exI) apply (rule_tac x=\' in exI) apply (rule_tac x=L' in exI) apply auto using E F apply simp apply (rule \) using ladder_E_L' apply blast using F finish apply blast using F finish apply blast by (rule L') qed theorem LeftDerivationLadder_cut_appendix: assumes LDLadder: "LeftDerivationLadder (\@\) D L \" assumes ladder_i_in_\: "ladder_i L 0 < length \" shows "\ E F \1 \2 L'. D = E@F \ \ = \1 @ \2 \ LeftDerivationLadder \ E L' \1 \ derivation_ge F (length \1) \ LeftDerivation \ (derivation_shift F (length \1) 0) \2 \ length L' = length L \ ladder_i L' 0 = ladder_i L 0 \ ladder_last_j L' = ladder_last_j L" proof - from LeftDerivationLadder_cut_appendix_helper[OF LDLadder ladder_i_in_\] obtain E F \1 \2 L' where helper: "D = E @ F \ \ = \1 @ \2 \ LeftDerivationLadder \ E L' \1 \ derivation_ge F (length \1) \ LeftDerivation \ (derivation_shift F (length \1) 0) \2 \ L' = ladder_cut L (length E)" by blast show ?thesis apply (rule_tac x=E in exI) apply (rule_tac x=F in exI) apply (rule_tac x=\1 in exI) apply (rule_tac x=\2 in exI) apply (rule_tac x=L' in exI) using helper LDLadder LeftDerivationLadder_def is_ladder_def ladder_i_of_cut_at_0 ladder_last_j_of_cut length_ladder_cut by force qed definition ladder_stepdown_diff :: "ladder \ nat" where "ladder_stepdown_diff L = Suc (ladder_n L 0)" definition ladder_stepdown_\_0 :: "sentence \ derivation \ ladder \ sentence" where "ladder_stepdown_\_0 a D L = Derive a (take (ladder_stepdown_diff L) D)" lemma LeftDerivationIntro_LeftDerives1: assumes "LeftDerivationIntro \ i r ix D j \" assumes "splits_at \ i a1 A a2" shows "LeftDerives1 \ i r (a1@(snd r)@a2)" by (metis LeftDerivationIntro_def LeftDerivationIntro_examine_rule LeftDerivationIntro_is_sentence LeftDerives1_def assms(1) assms(2) prod.collapse splits_at_implies_Derives1) lemma LeftDerives1_Derive: assumes "LeftDerives1 \ i r \" shows "Derive \ [(i, r)] = \" by (metis Derive LeftDerivation.simps(1) LeftDerivation_LeftDerives1 LeftDerivation_implies_Derivation append_Nil assms) lemma ladder_stepdown_\_0_altdef: assumes ladder: "LeftDerivationLadder \ D L \" assumes length_L: "length L > 1" assumes split: "splits_at (ladder_\ \ D L 1) (ladder_i L 1) a1 A a2" shows "ladder_stepdown_\_0 \ D L = a1 @ (snd (snd (D ! (ladder_n L 0)))) @ a2" proof - have 1: "ladder_\ \ D L 1 = Derive \ (take (ladder_n L 0) D)" by (simp add: ladder_\_def ladder_\_def) obtain rule where rule: "rule = snd (D ! (ladder_n L 0))" by blast have "\ E \. LeftDerivationIntro (ladder_\ \ D L 1) (ladder_i L 1) rule (ladder_ix L 1) E (ladder_j L 1) \" by (metis LeftDerivationIntrosAt_def LeftDerivationIntros_def LeftDerivationLadder_def One_nat_def diff_Suc_1 ladder length_L order_refl rule) then obtain E \ where intro: "LeftDerivationIntro (ladder_\ \ D L 1) (ladder_i L 1) rule (ladder_ix L 1) E (ladder_j L 1) \" by blast then have 2: "LeftDerives1 (ladder_\ \ D L 1) (ladder_i L 1) rule (a1@(snd rule)@a2)" using LeftDerivationIntro_LeftDerives1 split by blast have fst_D: "fst (D ! (ladder_n L 0)) = ladder_i L 1" by (metis LeftDerivationIntrosAt_def LeftDerivationIntros_def LeftDerivationLadder_def One_nat_def diff_Suc_1 ladder le_numeral_extra(4) length_L) have derive_derive: "Derive \ (take (Suc (ladder_n L 0)) D) = Derive (Derive \ (take (ladder_n L 0) D)) [D ! (ladder_n L 0)]" proof - have f1: "Derivation \ (take (Suc (ladder_n L 0)) D) (Derive \ (take (Suc (ladder_n L 0)) D))" using Derivation_take_derive LeftDerivationLadder_def LeftDerivation_implies_Derivation ladder by blast have f2: "length L - 1 < length L" using length_L by linarith have "0 < length L - 1" using length_L by linarith then have f3: "take (Suc (ladder_n L 0)) D = take (ladder_n L 0) D @ [D ! ladder_n L 0]" using f2 by (metis (full_types) LeftDerivationLadder_def is_ladder_def ladder ladder_last_n_def take_Suc_conv_app_nth) obtain sss :: "symbol list \ (nat \ symbol \ symbol list) list \ (nat \ symbol \ symbol list) list \ symbol list \ symbol list" where "\x0 x1 x2 x3. (\v4. Derivation x3 x2 v4 \ Derivation v4 x1 x0) = (Derivation x3 x2 (sss x0 x1 x2 x3) \ Derivation (sss x0 x1 x2 x3) x1 x0)" by moura then show ?thesis using f3 f1 Derivation_append Derive by auto qed then have 3: "ladder_stepdown_\_0 \ D L = Derive (ladder_\ \ D L 1) [D ! (ladder_n L 0)]" using 1 by (simp add: ladder_stepdown_\_0_def ladder_stepdown_diff_def) have 4: "D ! (ladder_n L 0) = (ladder_i L 1, rule)" using rule fst_D by (metis prod.collapse) then show ?thesis using 2 3 4 LeftDerives1_Derive snd_conv by auto qed lemma ladder_i_0_bound: assumes ld: "LeftDerivationLadder \ D L \" shows "ladder_i L 0 < length \" proof - have "LeftDerivationFix \ (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ \ D L 0)" using ld LeftDerivationLadder_def by simp then show ?thesis using LeftDerivationFix_def by simp qed lemma ladder_j_bound: assumes ld: "LeftDerivationLadder \ D L \" assumes index_bound: "index < length L" shows "ladder_j L index < length (ladder_\ \ D L index)" proof - have ld': "LeftDerivationLadder (\@[]) D L \" using ld by simp have ladder_i_0: "ladder_i L 0 < length \" using ladder_i_0_bound ld by auto obtain n where n: "n = ladder_n L index" by blast note propagate = LeftDerivationLadder_propagate[OF ld' ladder_i_0 n index_bound] from index_bound have "index + 1 < length L \ index + 1 = length L" by arith then show ?thesis proof (induct rule: disjCases2) case 1 then have "\\. LeftDerivation \ (take n D) \ \ ladder_\ (\ @ []) D L index = \ @ [] \ ladder_j L index < length \" using propagate by auto then show ?case by auto next case 2 then have " \n' \ \'. (index = 0 \ ladder_prev_n L index < n') \ n' \ n \ LeftDerivation \ (take n' D) \ \ LeftDerivation (\ @ []) (take n' D) (\ @ []) \ derivation_ge (drop n' D) (length \) \ LeftDerivation [] (derivation_shift (drop n' D) (length \) 0) \' \ ladder_\ (\ @ []) D L index = \ @ \' \ ladder_j L index < length \" using propagate by auto then show ?case by auto qed qed lemma ladder_last_j_bound: assumes ld: "LeftDerivationLadder \ D L \" shows "ladder_last_j L < length \" proof - have "length L - 1 < length L" using LeftDerivationLadder_def assms is_ladder_def by auto from ladder_j_bound[OF ld this] show ?thesis by (metis Derive LeftDerivationLadder_def LeftDerivation_implies_Derivation One_nat_def is_ladder_def ladder_last_j_def last_ladder_\ ld) qed fun ladder_shift_n :: "nat \ ladder \ ladder" where "ladder_shift_n N [] = []" | "ladder_shift_n N ((n, j, i)#L) = ((n - N, j, i)#(ladder_shift_n N L))" fun ladder_stepdown :: "ladder \ ladder" where "ladder_stepdown [] = undefined" | "ladder_stepdown [v] = undefined" | "ladder_stepdown ((n0, j0, i0)#(n1, j1, ix1)#L) = (n1 - Suc n0, j1, j0 + ix1) # (ladder_shift_n (Suc n0) L)" lemma ladder_shift_n_length: "length (ladder_shift_n N L) = length L" by (induct L, auto) lemma ladder_stepdown_prepare: assumes "length L > 1" shows "L = (ladder_n L 0, ladder_j L 0, ladder_i L 0)# (ladder_n L 1, ladder_j L 1, ladder_ix L 1)#(drop 2 L)" proof - have "\ n0 j0 i0 n1 j1 ix1 L'. L = ((n0, j0, i0)#(n1, j1, ix1)#L')" by (metis One_nat_def Suc_eq_plus1 assms ladder_stepdown.cases less_not_refl list.size(3) list.size(4) not_less0) then obtain n0 j0 i0 n1 j1 ix1 L' where L': "L = ((n0, j0, i0)#(n1, j1, ix1)#L')" by blast have n0: "n0 = ladder_n L 0" using L' by (auto simp add: ladder_n_def deriv_n_def) show ?thesis using L' by (auto simp add: ladder_n_def deriv_n_def ladder_j_def deriv_j_def ladder_i_def deriv_i_def ladder_ix_def deriv_ix_def) qed lemma ladder_stepdown_length: assumes "length L > 1" shows "length (ladder_stepdown L) = length L - 1" apply (subst ladder_stepdown_prepare[OF assms(1)]) apply (simp add: ladder_shift_n_length) using assms apply arith done lemma ladder_stepdown_i_0: assumes "length L > 1" shows "ladder_i (ladder_stepdown L) 0 = ladder_i L 1 + ladder_ix L 1" using ladder_stepdown_prepare[OF assms] ladder_i_def ladder_j_def deriv_j_def by (metis One_nat_def deriv_i_def diff_Suc_1 ladder_stepdown.simps(3) list.sel(1) snd_conv zero_neq_one) lemma ladder_shift_n_cons: "ladder_shift_n N (x#L) = (fst x - N, snd x)#(ladder_shift_n N L)" apply (induct L) by (cases x, simp)+ lemma ladder_shift_n_drop: "ladder_shift_n N (drop n L) = drop n (ladder_shift_n N L)" proof (induct L arbitrary: n) case Nil then show ?case by simp next case (Cons x L) show ?case proof (cases n) case 0 then show ?thesis by simp next case (Suc n) then show ?thesis by (simp add: ladder_shift_n_cons Cons) qed qed lemma drop_2_shift: assumes "index > 0" assumes "length L > 1" shows "drop 2 L ! (index - Suc 0) = L ! Suc index" proof - define l1 l2 and L' where "l1 = L ! 0" "l2 = L ! 1" and "L' = drop 2 L" with \length L > 1\ have "L = l1 # l2 # L'" by (cases L) (auto simp add: neq_Nil_conv) with \index > 0\ show ?thesis by simp qed lemma ladder_shift_n_at: "index < length L \ (ladder_shift_n N L) ! index = (fst (L ! index) - N, snd (L ! index))" proof (induct L arbitrary: index) case Nil then show ?case by auto next case (Cons x L) show ?case apply (simp add: ladder_shift_n_cons) apply (cases index) apply (auto) apply (rule_tac Cons(1)) using Cons(2) by auto qed lemma ladder_stepdown_j: assumes length_L_greater_1: "length L > 1" assumes L': "L' = ladder_stepdown L" assumes index_bound: "index < length L'" shows "ladder_j L' index = ladder_j L (Suc index)" proof - note L_prepare = ladder_stepdown_prepare[OF length_L_greater_1] have ladder_stepdown_L_def: "ladder_stepdown L = ((ladder_n L (Suc 0) - Suc (ladder_n L 0), ladder_j L (Suc 0), ladder_j L 0 + ladder_ix L (Suc 0)) # ladder_shift_n (Suc (ladder_n L 0)) (drop 2 L))" by (subst L_prepare, simp) have "index = 0 \ index > 0" by arith then show "ladder_j L' index = ladder_j L (Suc index)" proof (induct rule: disjCases2) case 1 show ?case by (simp add: L' ladder_stepdown_L_def 1 ladder_j_def deriv_j_def) next case 2 have index_bound': "Suc index < length L" using index_bound L' ladder_stepdown_length length_L_greater_1 by auto show ?case apply (simp add: L' ladder_stepdown_L_def 2 ladder_j_def ladder_shift_n_drop drop_2_shift) apply (subst drop_2_shift) apply (simp add: 2) using length_L_greater_1 apply (simp add: ladder_shift_n_length) apply (simp add: deriv_j_def) apply (simp add: ladder_shift_n_at[OF index_bound']) done qed qed lemma ladder_stepdown_last_j: assumes length_L_greater_1: "length L > 1" shows "ladder_last_j (ladder_stepdown L) = ladder_last_j L" using ladder_stepdown_j Suc_diff_Suc diff_Suc_1 ladder_last_j_def ladder_stepdown_length length_L_greater_1 lessI by auto lemma ladder_stepdown_n: assumes length_L_greater_1: "length L > 1" assumes L': "L' = ladder_stepdown L" assumes index_bound: "index < length L'" shows "ladder_n L' index = ladder_n L (Suc index) - ladder_stepdown_diff L" proof - note L_prepare = ladder_stepdown_prepare[OF length_L_greater_1] have ladder_stepdown_L_def: "ladder_stepdown L = ((ladder_n L (Suc 0) - Suc (ladder_n L 0), ladder_j L (Suc 0), ladder_j L 0 + ladder_ix L (Suc 0)) # ladder_shift_n (Suc (ladder_n L 0)) (drop 2 L))" by (subst L_prepare, simp) have "index = 0 \ index > 0" by arith then show "ladder_n L' index = ladder_n L (Suc index) - ladder_stepdown_diff L" proof (induct rule: disjCases2) case 1 show ?case by (simp add: L' ladder_stepdown_L_def 1 ladder_n_def deriv_n_def ladder_stepdown_diff_def) next case 2 have index_bound': "Suc index < length L" using index_bound L' ladder_stepdown_length length_L_greater_1 by auto show ?case apply (simp add: L' ladder_stepdown_L_def 2 ladder_n_def ladder_shift_n_drop drop_2_shift ladder_stepdown_diff_def) apply (subst drop_2_shift) apply (simp add: 2) using length_L_greater_1 apply (simp add: ladder_shift_n_length) apply (simp add: deriv_n_def) apply (simp add: ladder_shift_n_at[OF index_bound']) done qed qed lemma ladder_stepdown_ix: assumes length_L_greater_1: "length L > 1" assumes L': "L' = ladder_stepdown L" assumes index_lower_bound: "0 < index" assumes index_upper_bound: "index < length L'" shows "ladder_ix L' index = ladder_ix L (Suc index)" proof - note L_prepare = ladder_stepdown_prepare[OF length_L_greater_1] have ladder_stepdown_L_def: "ladder_stepdown L = ((ladder_n L (Suc 0) - Suc (ladder_n L 0), ladder_j L (Suc 0), ladder_j L 0 + ladder_ix L (Suc 0)) # ladder_shift_n (Suc (ladder_n L 0)) (drop 2 L))" by (subst L_prepare, simp) have index_bound': "Suc index < length L" using index_upper_bound L' ladder_stepdown_length length_L_greater_1 by auto show ?thesis apply (simp add: L' ladder_stepdown_L_def index_lower_bound ladder_ix_def ladder_shift_n_drop) apply (subst drop_2_shift) apply (simp add: index_lower_bound) using length_L_greater_1 apply (simp add: ladder_shift_n_length) apply (simp add: deriv_ix_def) apply (simp add: ladder_shift_n_at[OF index_bound']) using index_lower_bound by arith qed lemma Derive_Derive: assumes "Derivation \ (D@E) \" shows "Derive (Derive \ D) E = Derive \ (D@E)" using Derivation_append Derive assms by fastforce lemma drop_at_shift: assumes "n \ index" assumes "index < length D" shows "drop n D ! (index - n) = D ! index" using assms(1) assms(2) by auto theorem LeftDerivationLadder_stepdown: assumes ldl: "LeftDerivationLadder \ D L \" assumes length_L: "length L > 1" shows "\ L'. LeftDerivationLadder (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' \ \ length L' = length L - 1 \ ladder_i L' 0 = ladder_i L 1 + ladder_ix L 1 \ ladder_last_j L' = ladder_last_j L" proof - obtain L' where L': "L' = ladder_stepdown L" by blast have ldl1: "LeftDerivation (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) \" proof - have D_split: "D = (take (ladder_stepdown_diff L) D) @ (drop (ladder_stepdown_diff L) D)" by simp show ?thesis using D_split ldl proof - obtain sss :: "symbol list \ (nat \ symbol \ symbol list) list \ (nat \ symbol \ symbol list) list \ symbol list \ symbol list" where "\x0 x1 x2 x3. (\v4. LeftDerivation x3 x2 v4 \ LeftDerivation v4 x1 x0) = (LeftDerivation x3 x2 (sss x0 x1 x2 x3) \ LeftDerivation (sss x0 x1 x2 x3) x1 x0)" by moura then have "(\ LeftDerivation \ (take (ladder_stepdown_diff L) D @ drop (ladder_stepdown_diff L) D) \ \ LeftDerivation \ (take (ladder_stepdown_diff L) D) (sss \ (drop (ladder_stepdown_diff L) D) (take (ladder_stepdown_diff L) D) \) \ LeftDerivation (sss \ (drop (ladder_stepdown_diff L) D) (take (ladder_stepdown_diff L) D) \) (drop (ladder_stepdown_diff L) D) \) \ (LeftDerivation \ (take (ladder_stepdown_diff L) D @ drop (ladder_stepdown_diff L) D) \ \ (\ss. \ LeftDerivation \ (take (ladder_stepdown_diff L) D) ss \ \ LeftDerivation ss (drop (ladder_stepdown_diff L) D) \))" using LeftDerivation_append by blast then show ?thesis by (metis (no_types) D_split Derivation_take_derive Derivation_unique_dest LeftDerivationLadder_def LeftDerivation_implies_Derivation ladder_stepdown_\_0_def ldl) qed qed have L'_nonempty: "L' \ []" using L' ladder_stepdown_length length_L by fastforce { fix u :: nat assume u': "u < length L'" then have Suc_u: "Suc u < length L" using L' ladder_stepdown_length length_L by auto have "ladder_n L (Suc u) \ length D" using ldl Suc_u by (simp add: LeftDerivationLadder_ladder_n_bound) then have "ladder_n L' u \ length D - ladder_stepdown_diff L" apply (subst ladder_stepdown_n[OF length_L L' u']) by auto } note is_ladder_prop1 = this { fix u :: nat fix v :: nat assume u_less_v: "u < v" assume v_L': "v < length L'" from u_less_v v_L' have u_L': "u < length L'" by arith have "ladder_n L (Suc u) < ladder_n L (Suc v)" using ldl by (metis (no_types, lifting) L' LeftDerivationLadder_def One_nat_def Suc_diff_1 Suc_lessD Suc_mono is_ladder_def ladder_stepdown_length length_L u_less_v v_L') then have "ladder_n L' u < ladder_n L' v" apply (simp add: ladder_stepdown_n[OF length_L L'] u_L' v_L') by (metis (no_types, lifting) L' LeftDerivationLadder_def Suc_eq_plus1 Suc_leI diff_less_mono is_ladder_def ladder_stepdown_diff_def ladder_stepdown_length ldl length_L less_diff_conv u_L' zero_less_Suc) } note is_ladder_prop2 = this have is_ladder_L': "is_ladder (drop (ladder_stepdown_diff L) D) L'" apply (auto simp add: is_ladder_def) using L'_nonempty apply blast using is_ladder_prop1 apply blast using is_ladder_prop2 apply blast using ladder_last_n_def ladder_stepdown_n L' LeftDerivationLadder_def Suc_diff_Suc diff_Suc_1 ladder_n_last_is_length ladder_stepdown_length ldl length_L lessI by auto have ldfix: "LeftDerivationFix (ladder_stepdown_\_0 \ D L) (ladder_i L' 0) (take (ladder_n L' 0) (drop (ladder_stepdown_diff L) D)) (ladder_j L' 0) (ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' 0)" proof - have introsAt_L_1: "LeftDerivationIntrosAt \ D L 1" using LeftDerivationIntros_def LeftDerivationLadder_def ldl length_L by blast thm LeftDerivationIntrosAt_def obtain n where n: "n = ladder_n L 0" by blast obtain m where m: "m = ladder_n L 1" by blast have "LeftDerivationIntro (ladder_\ \ D L 1) (ladder_i L 1) (snd (D ! n)) (ladder_ix L 1) (drop (Suc n) (take m D)) (ladder_j L 1) (ladder_\ \ D L 1)" using n m introsAt_L_1 by (metis LeftDerivationIntrosAt_def One_nat_def diff_Suc_1) from iffD1[OF LeftDerivationIntro_def this] obtain \ where \: "LeftDerives1 (ladder_\ \ D L 1) (ladder_i L 1) (snd (D ! n)) \ \ ladder_ix L 1 < length (snd (snd (D ! n))) \ snd (snd (D ! n)) ! ladder_ix L 1 = ladder_\ \ D L 1 ! ladder_j L 1 \ LeftDerivationFix \ (ladder_i L 1 + ladder_ix L 1) (drop (Suc n) (take m D)) (ladder_j L 1) (ladder_\ \ D L 1)" by blast have "\ = Derive (ladder_\ \ D L 1) [D ! n]" by (metis (no_types, hide_lams) LeftDerivationIntrosAt_def LeftDerives1_Derive \ cancel_comm_monoid_add_class.diff_cancel introsAt_L_1 n prod.collapse) then have \_def: "\ = ladder_stepdown_\_0 \ D L" proof - obtain sss :: "nat \ symbol list \ symbol list" and ss :: "nat \ symbol list \ symbol" and sssa :: "nat \ symbol list \ symbol list" where "\x2 x3. (\v4 v5 v6. splits_at x3 x2 v4 v5 v6) = splits_at x3 x2 (sss x2 x3) (ss x2 x3) (sssa x2 x3)" by moura then have f1: "\ssa n p ssb. \ Derives1 ssa n p ssb \ splits_at ssa n (sss n ssa) (ss n ssa) (sssa n ssa)" using splits_at_ex by presburger then have "\ = sss (ladder_i L 1) (ladder_\ \ D L 1) @ snd (snd (D ! n)) @ sssa (ladder_i L 1) (ladder_\ \ D L 1)" by (meson LeftDerives1_implies_Derives1 \ splits_at_combine_dest) then show ?thesis using f1 by (metis (no_types) LeftDerives1_implies_Derives1 \ ladder_stepdown_\_0_altdef ldl length_L n) qed have ladder_i_L'_0: "ladder_i L' 0 = ladder_i L 1 + ladder_ix L 1" using L' ladder_stepdown_i_0 length_L by blast have derivation_eq: "(take (ladder_n L' 0) (drop (ladder_stepdown_diff L) D)) = (drop (Suc n) (take m D))" using n m by (metis L' L'_nonempty One_nat_def drop_take ladder_stepdown_diff_def ladder_stepdown_n length_L length_greater_0_conv) have ladder_j_L'_0: "ladder_j L' 0 = ladder_j L 1" using L' L'_nonempty ladder_stepdown_j length_L by auto have ladder_\: "(ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' 0) = ladder_\ \ D L 1" by (metis Derivation_take_derive Derivation_unique_dest LeftDerivationFix_def LeftDerivation_implies_Derivation \ \_def derivation_eq ladder_\_def ldl1) from \_def \ ladder_i_L'_0 derivation_eq ladder_j_L'_0 ladder_\ show ?thesis by auto qed { fix index :: nat assume index_lower_bound: "Suc 0 \ index" assume index_upper_bound: "index < length L'" then have Suc_index_upper_bound: "Suc index < length L" using L' Suc_diff_Suc Suc_less_eq diff_Suc_1 ladder_stepdown_length length_L less_Suc_eq by auto then have intros_at_Suc_index: "LeftDerivationIntrosAt \ D L (Suc index)" by (metis LeftDerivationIntros_def LeftDerivationLadder_def Suc_eq_plus1_left ldl le_add1) from iffD1[OF LeftDerivationIntrosAt_def this] have ldintro: "let \' = ladder_\ \ D L (Suc index); i = ladder_i L (Suc index); j = ladder_j L (Suc index); ix = ladder_ix L (Suc index); \ = ladder_\ \ D L (Suc index); n = ladder_n L (Suc index - 1); m = ladder_n L (Suc index); e = D ! n; E = drop (Suc n) (take m D) in i = fst e \ LeftDerivationIntro \' i (snd e) ix E j \" by blast have index_minus_Suc_0_bound: "index - Suc 0 < length L'" by (simp add: index_upper_bound less_imp_diff_less) note helpers = length_L L' index_minus_Suc_0_bound have ladder_i_L'_index: "ladder_i L' index = fst (drop (ladder_stepdown_diff L) D ! ladder_n L' (index - Suc 0))" apply (auto simp add: ladder_i_def) using index_lower_bound apply arith apply (simp add: ladder_stepdown_n[OF helpers] ladder_stepdown_j[OF helpers]) apply (subst drop_at_shift) using LeftDerivationLadder_def Suc_index_upper_bound Suc_leI Suc_lessD is_ladder_def ladder_stepdown_diff_def ldl apply presburger apply (metis LeftDerivationLadder_def One_nat_def Suc_eq_plus1 Suc_index_upper_bound add.commute add_diff_cancel_right' ladder_n_minus_1_bound ldl le_add1) by (metis LeftDerivationIntrosAt_def intros_at_Suc_index diff_Suc_1 ladder_i_def nat.simps(3)) have intro_at_index: "LeftDerivationIntro (ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' index) (ladder_i L' index) (snd (drop (ladder_stepdown_diff L) D ! ladder_n L' (index - Suc 0))) (ladder_ix L' index) (drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) (drop (ladder_stepdown_diff L) D))) (ladder_j L' index) (ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' index)" proof - have arg1: "(ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' index) = ladder_\ \ D L (Suc index)" apply (auto simp add: ladder_\_def ladder_\_def) using index_lower_bound apply arith apply (simp add: ladder_stepdown_n[OF helpers] ladder_stepdown_\_0_def) apply (subst Derive_Derive[where \="ladder_\ \ D L index"]) apply (metis (no_types, lifting) Derivation_take_derive LeftDerivationLadder_def LeftDerivation_implies_Derivation Suc_index_upper_bound Suc_leI Suc_lessD add.commute is_ladder_def ladder_\_def ladder_stepdown_diff_def ldl le_add_diff_inverse2 take_add) by (metis LeftDerivationLadder_def Suc_index_upper_bound Suc_leI Suc_lessD add.commute is_ladder_def ladder_stepdown_diff_def ldl le_add_diff_inverse2 take_add) have arg2: "ladder_i L' index = ladder_i L (Suc index)" using L' index_lower_bound index_minus_Suc_0_bound ladder_i_def ladder_stepdown_j length_L by auto obtain n where n: "n = ladder_n L (Suc index - 1)" by blast have arg3: "(snd (drop (ladder_stepdown_diff L) D ! ladder_n L' (index - Suc 0))) = snd (D ! n)" apply (simp add: ladder_stepdown_n[OF helpers] index_lower_bound) apply (subst drop_at_shift) using index_lower_bound apply (metis (no_types, hide_lams) L' LeftDerivationLadder_def One_nat_def Suc_eq_plus1 add.commute diff_Suc_1 index_upper_bound is_ladder_def ladder_stepdown_diff_def ladder_stepdown_length ldl le_add_diff_inverse2 length_L less_or_eq_imp_le n nat.simps(3) neq0_conv not_less not_less_eq_eq) using index_lower_bound apply (metis LeftDerivationLadder_def One_nat_def Suc_index_upper_bound Suc_le_lessD Suc_pred diff_Suc_1 ladder_n_minus_1_bound ldl le_imp_less_Suc less_imp_le) using index_lower_bound n by (simp add: Suc_diff_le) have arg4: "ladder_ix L' index = ladder_ix L (Suc index)" using ladder_stepdown_ix L' Suc_le_lessD index_lower_bound index_upper_bound length_L by auto obtain m where m: "m = ladder_n L (Suc index)" by blast have Suc_index_Suc: "Suc (index - Suc 0) = index" using index_lower_bound by arith have arg5: "(drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) (drop (ladder_stepdown_diff L) D))) = drop (Suc n) (take m D)" apply (simp add: ladder_stepdown_n[OF helpers] ladder_stepdown_n[OF length_L L' index_upper_bound] n m Suc_index_Suc) by (metis (no_types, lifting) LeftDerivationLadder_def Suc_eq_plus1_left Suc_index_upper_bound Suc_leI Suc_le_lessD Suc_lessD drop_drop drop_take index_lower_bound is_ladder_def ladder_stepdown_diff_def ldl le_add_diff_inverse2) have arg6: "ladder_j L' index = ladder_j L (Suc index)" using L' index_upper_bound ladder_stepdown_j length_L by blast have arg7: "(ladder_\ (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' index) = ladder_\ \ D L (Suc index)" apply (simp add: ladder_\_def) apply (simp add: ladder_stepdown_n[OF length_L L' index_upper_bound] ladder_stepdown_\_0_def) apply (subst Derive_Derive[where \="ladder_\ \ D L (Suc index)"]) apply (metis (no_types, lifting) L' LeftDerivationLadder_def LeftDerivation_implies_Derivation LeftDerivation_take_derive Suc_le_lessD add_diff_inverse_nat diff_is_0_eq index_lower_bound index_upper_bound is_ladder_L' is_ladder_def ladder_\_def ladder_stepdown_n ldl le_0_eq length_L less_numeral_extra(3) less_or_eq_imp_le take_add) by (metis (no_types, lifting) L' One_nat_def add_diff_inverse_nat diff_is_0_eq index_lower_bound index_upper_bound is_ladder_L' is_ladder_def ladder_stepdown_n le_0_eq le_neq_implies_less length_L less_numeral_extra(3) less_or_eq_imp_le take_add zero_less_one) from ldintro arg1 arg2 arg3 arg4 arg5 arg6 arg7 show ?thesis by (metis m n) qed have "LeftDerivationIntrosAt (ladder_stepdown_\_0 \ D L) (drop (ladder_stepdown_diff L) D) L' index" apply (auto simp add: LeftDerivationIntrosAt_def Let_def) using ladder_i_L'_index apply blast using intro_at_index by blast } note introsAt = this show ?thesis apply (rule_tac x="L'" in exI) apply auto defer 1 using L' ladder_stepdown_length length_L apply auto[1] using ladder_stepdown_i_0 length_L L' apply auto[1] using ladder_stepdown_last_j L' length_L apply auto[1] apply (auto simp add: LeftDerivationLadder_def) using ldl1 apply blast using is_ladder_L' apply blast using ldfix apply blast apply (auto simp add: LeftDerivationIntros_def) apply (simp add: introsAt) done qed fun ladder_shift_j :: "nat \ ladder \ ladder" where "ladder_shift_j d [] = []" | "ladder_shift_j d ((n, j, i)#L) = ((n, j - d, i)#(ladder_shift_j d L))" definition ladder_cut_prefix :: "nat \ ladder \ ladder" where "ladder_cut_prefix d L = (ladder_shift_j d L)[0 := (ladder_n L 0, ladder_j L 0 - d, ladder_i L 0 - d)]" lemma ladder_shift_j_length: "length (ladder_shift_j d L) = length L" by (induct L, auto) lemma ladder_cut_prefix_length: shows "length (ladder_cut_prefix d L) = length L" apply (simp add: ladder_cut_prefix_def) apply (simp add: ladder_shift_j_length) done lemma ladder_shift_j_cons: "ladder_shift_j d (x#L) = (fst x, fst (snd x) - d, snd(snd x))# (ladder_shift_j d L)" apply (induct L) by (cases x, simp)+ lemma deriv_j_ladder_shift_j: "index < length L \ deriv_j (ladder_shift_j d L ! index) = deriv_j (L ! index) - d" proof (induct L arbitrary: index) case Nil then show ?case by auto next case (Cons x L) show ?case apply (subst ladder_shift_j_cons) apply (cases index) using Cons by (auto simp add: deriv_j_def) qed lemma deriv_n_ladder_shift_j: "index < length L \ deriv_n (ladder_shift_j d L ! index) = deriv_n (L ! index)" proof (induct L arbitrary: index) case Nil then show ?case by auto next case (Cons x L) show ?case apply (subst ladder_shift_j_cons) apply (cases index) using Cons by (auto simp add: deriv_n_def) qed lemma deriv_ix_ladder_shift_j: "index < length L \ deriv_ix (ladder_shift_j d L ! index) = deriv_ix (L ! index)" proof (induct L arbitrary: index) case Nil then show ?case by auto next case (Cons x L) show ?case apply (subst ladder_shift_j_cons) apply (cases index) using Cons by (auto simp add: deriv_ix_def) qed lemma ladder_cut_prefix_j: assumes index_bound: "index < length L" assumes length_L: "length L > 0" shows "ladder_j (ladder_cut_prefix d L) index = ladder_j L index - d" apply (simp add: ladder_j_def ladder_cut_prefix_def) apply (cases index) apply (auto simp add: length_L) apply (subst nth_list_update_eq) apply (simp only: ladder_shift_j_length length_L) apply (simp add: deriv_j_def) apply (subst deriv_j_ladder_shift_j) using index_bound apply arith by blast lemma hd_0_subst: "length L > 0 \ hd (L [0 := x]) = x" using hd_conv_nth by (simp add: upd_conv_take_nth_drop) lemma ladder_cut_prefix_i: assumes index_bound: "index < length L" assumes length_L: "length L > 0" shows "ladder_i (ladder_cut_prefix d L) index = ladder_i L index - d" apply (simp add: ladder_i_def ladder_cut_prefix_def) apply (cases index) apply auto[1] apply (subst hd_0_subst) using length_L ladder_shift_j_length apply metis apply (auto simp add: deriv_i_def) apply (case_tac nat) apply (simp add: ladder_j_def deriv_j_def) apply auto apply (subst nth_list_update_eq) using length_L ladder_shift_j_length apply auto[1] apply simp apply (simp add: ladder_j_def) apply (subst deriv_j_ladder_shift_j) using index_bound apply arith apply simp done lemma ladder_cut_prefix_n: assumes index_bound: "index < length L" assumes length_L: "length L > 0" shows "ladder_n (ladder_cut_prefix d L) index = ladder_n L index" apply (simp add: ladder_cut_prefix_def) apply (cases index) apply (auto simp add: ladder_n_def) apply (subst nth_list_update_eq) apply (simp add: ladder_shift_j_length) using length_L apply blast apply (simp add: deriv_n_def ) apply (rule_tac deriv_n_ladder_shift_j) using index_bound by arith lemma ladder_cut_prefix_ix: assumes index_bound: "index < length L" assumes length_L: "length L > 0" shows "ladder_ix (ladder_cut_prefix d L) index = ladder_ix L index" apply (simp add: ladder_cut_prefix_def) apply (cases index) apply (auto simp add: ladder_ix_def) apply (rule_tac deriv_ix_ladder_shift_j) using index_bound by arith lemma LeftDerivationFix_derivation_ge_is_nonterminal: assumes ldfix: "LeftDerivationFix \ i D j \" assumes derivation_ge_d: "derivation_ge D d" assumes is_nonterminal: "is_nonterminal (\ ! j)" shows "(D = [] \ \ = \ \ i = j) \ (i > d \ j \ d)" proof - have "is_nonterminal (\ ! i)" using ldfix is_nonterminal by (simp add: LeftDerivationFix_def) from LeftDerivationFix_splits_at_nonterminal[OF ldfix this] obtain U a1 a2 b1 where U: "splits_at \ i a1 U a2 \ splits_at \ j b1 U a2 \ LeftDerivation a1 D b1" by blast have "D = [] \ D \ []" by auto then show ?thesis proof (induct rule: disjCases2) case 1 then have "a1 = b1" using U by auto then have i_eq_j: "i = j" using U by (metis dual_order.strict_implies_order length_take min.absorb2 splits_at_def) from 1 have "\ = \" using ldfix LeftDerivationFix_def by auto with 1 i_eq_j show ?case by blast next case 2 have "\ a1'. LeftDerives1 a1 (fst (hd D)) (snd (hd D)) a1'" using U 2 by (metis LeftDerivation.elims(2) list.sel(1)) then obtain a1' where a1': "LeftDerives1 a1 (fst (hd D)) (snd (hd D)) a1'" by blast then have "(fst (hd D)) < length a1" using Derives1_bound LeftDerives1_implies_Derives1 by blast then have fst_less_i: "(fst (hd D)) < i" using U by (simp add: leD min.absorb2 nat_le_linear splits_at_def) have d_le_fst: "d \ fst (hd D)" using derivation_ge_d 2 by (simp add: derivation_ge_def) with fst_less_i have d_less_i: "d < i" using le_less_trans by blast have "\ b1'. LeftDerives1 b1' (fst (last D)) (snd (last D)) b1" using U 2 by (metis Derive LeftDerivation_Derive_take_LeftDerives1 LeftDerivation_implies_Derivation last_conv_nth length_0_conv order_refl take_all) then obtain b1' where b1': "LeftDerives1 b1' (fst (last D)) (snd (last D)) b1" by blast then have "fst (last D) \ length b1" using Derives1_bound' LeftDerives1_implies_Derives1 by blast then have fst_le_j: "fst (last D) \ j" using U splits_at_def by auto have "d \ fst (last D)" using derivation_ge_d 2 using derivation_ge_def last_in_set by blast with fst_le_j have "d \ j" using order.trans by blast with d_less_i show ?thesis by auto qed qed lemma LeftDerivationFix_derivation_ge: assumes ldfix: "LeftDerivationFix \ i D j \" assumes derivation_ge_d: "derivation_ge D d" shows "i = j \ (i > d \ j \ d)" proof - from LeftDerivationFix_splits_at_symbol[OF ldfix] obtain U a1 a2 b1 b2 n where U: "splits_at \ i a1 U a2 \ splits_at \ j b1 U b2 \ n \ length D \ LeftDerivation a1 (take n D) b1 \ derivation_ge (drop n D) (Suc (length b1)) \ LeftDerivation a2 (derivation_shift (drop n D) (Suc (length b1)) 0) b2 \ (n = length D \ n < length D \ is_word (b1 @ [U]))" by blast have "n = 0 \ n > 0" by auto then show ?thesis proof (induct rule: disjCases2) case 1 then have "a1 = b1" using U by auto then have i_eq_j: "i = j" using U by (metis dual_order.strict_implies_order length_take min.absorb2 splits_at_def) then show ?case by blast next case 2 obtain E where E: "E = take n D" by blast have E_nonempty: "E \ []" using E 2 using U less_nat_zero_code list.size(3) take_eq_Nil by auto have "\ a1'. LeftDerives1 a1 (fst (hd E)) (snd (hd E)) a1'" using U E E_nonempty by (metis LeftDerivation.simps(2) list.exhaust list.sel(1)) then obtain a1' where a1': "LeftDerives1 a1 (fst (hd E)) (snd (hd E)) a1'" by blast then have "(fst (hd E)) < length a1" using Derives1_bound LeftDerives1_implies_Derives1 by blast then have fst_less_i: "(fst (hd E)) < i" using U by (simp add: leD min.absorb2 nat_le_linear splits_at_def) have d_le_fst: "d \ fst (hd E)" using derivation_ge_d E_nonempty E by (simp add: LeftDerivation.elims(2) U derivation_ge_def hd_conv_nth) with fst_less_i have d_less_i: "d < i" using le_less_trans by blast have "\ b1'. LeftDerives1 b1' (fst (last E)) (snd (last E)) b1" using E_nonempty U E by (metis LeftDerivation_append1 append_butlast_last_id prod.collapse) then obtain b1' where b1': "LeftDerives1 b1' (fst (last E)) (snd (last E)) b1" by blast then have "fst (last E) \ length b1" using Derives1_bound' LeftDerives1_implies_Derives1 by blast then have fst_le_j: "fst (last E) \ j" using U splits_at_def by auto have "d \ fst (last E)" using derivation_ge_d E_nonempty E using derivation_ge_d last_in_set by (metis derivation_ge_def set_take_subset subsetCE) with fst_le_j have "d \ j" using order.trans by blast with d_less_i show ?thesis by auto qed qed lemma LeftDerivationIntro_derivation_ge: assumes ldintro: "LeftDerivationIntro \ i r ix D j \" assumes i_ge_d: "i \ d" assumes derivation_ge_d: "derivation_ge D d" shows "j \ d" proof - from iffD1[OF LeftDerivationIntro_def ldintro] obtain \ where \: "LeftDerives1 \ i r \ \ ix < length (snd r) \ snd r ! ix = \ ! j \ LeftDerivationFix \ (i + ix) D j \" by blast then have "(i + ix = j) \ (i + ix > d \ j \ d)" using LeftDerivationFix_derivation_ge derivation_ge_d by blast then show ?thesis proof (induct rule: disjCases2) case 1 then show ?case using i_ge_d trans_le_add1 by blast next case 2 then show ?case by simp qed qed lemma derivation_ge_LeftDerivationLadder: assumes derivation_ge_d: "derivation_ge D d" assumes ladder: "LeftDerivationLadder \ D L \" assumes ladder_i_0: "ladder_i L 0 \ d" shows "index < length L \ ladder_i L index \ d \ ladder_j L index \ d" proof (induct index) case 0 from iffD1[OF LeftDerivationLadder_def ladder] have ldfix: "LeftDerivationFix \ (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ \ D L 0)" by blast have "derivation_ge (take (ladder_n L 0) D) d" using derivation_ge_d by (metis append_take_drop_id derivation_ge_append) from ladder_i_0 derivation_ge_d LeftDerivationFix_derivation_ge[OF ldfix this] show ?case by linarith next case (Suc n) have ladder_i_Suc: "ladder_i L (Suc n) \ d" apply (auto simp add: ladder_i_def) using Suc by auto from iffD1[OF LeftDerivationLadder_def ladder] have "LeftDerivationIntros \ D L" by blast then have "LeftDerivationIntrosAt \ D L (Suc n)" using Suc.prems by (metis LeftDerivationIntros_def Suc_eq_plus1_left le_add1) from iffD1[OF LeftDerivationIntrosAt_def this] show ?case using ladder_i_Suc LeftDerivationIntro_derivation_ge by (metis append_take_drop_id derivation_ge_append derivation_ge_d) qed lemma derivation_shift_append: "derivation_shift (A@B) left right = (derivation_shift A left right) @ (derivation_shift B left right)" by (induct A, simp+) lemma derivation_shift_right_left_subtract: "right \ left \ derivation_shift (derivation_shift L 0 right) left 0 = derivation_shift L 0 (right - left)" by (induct L, simp+) lemma LeftDerivationFix_cut_prefix: assumes "LeftDerivationFix (\@\) i D j \" assumes "derivation_ge D (length \)" assumes "i \ length \" assumes is_word_\: "is_word \" shows "\ \'. \ = \ @ \' \ LeftDerivationFix \ (i - length \) (derivation_shift D (length \) 0) (j - length \) \'" proof - have j_ge_d: "j \ length \" using assms(3) LeftDerivationFix_derivation_ge[OF assms(1) assms(2)] by arith obtain \' where \': "\' = drop (length \) \" by blast from iffD1[OF LeftDerivationFix_def assms(1)] obtain E F where EF: "is_sentence (\ @ \) \ is_sentence \ \ LeftDerivation (\ @ \) D \ \ i < length (\ @ \) \ j < length \ \ (\ @ \) ! i = \ ! j \ D = E @ derivation_shift F 0 (Suc j) \ LeftDerivation (take i (\ @ \)) E (take j \) \ LeftDerivation (drop (Suc i) (\ @ \)) F (drop (Suc j) \)" by blast then have "LeftDerivation (\ @ \) D \" by blast from LeftDerivation_skip_prefixword_ex[OF this is_word_\] obtain \' where \': "\ = \ @ \' \ LeftDerivation \ (derivation_shift D (length \) 0) \'" by blast have ldf1: "is_sentence \" using EF is_sentence_concat by blast have ldf2: "is_sentence \'" using EF \' is_sentence_concat by blast have ldf3: "i - length \ < length \" by (metis EF append_Nil assms(3) drop_append drop_eq_Nil not_le) have ldf4: "j - length \ < length \'" by (metis EF append_Nil j_ge_d \' drop_append drop_eq_Nil not_le) have ldf5: "\ ! (i - length \) = \' ! (j - length \)" by (metis \' EF assms(3) j_ge_d leD nth_append) have D_split: "D = E @ derivation_shift F 0 (Suc j)" using EF by blast show ?thesis apply (rule_tac x=\' in exI) apply (auto simp add: \') apply (auto simp add: LeftDerivationFix_def) using ldf1 apply blast using ldf2 apply blast using \' apply blast using ldf3 apply blast using ldf4 apply blast using ldf5 apply blast apply (rule_tac x="derivation_shift E (length \) 0" in exI) apply (rule_tac x="F" in exI) apply auto apply (subst D_split) apply (simp add: derivation_shift_append) apply (subst derivation_shift_right_left_subtract) apply (simp add: j_ge_d le_Suc_eq) using j_ge_d apply (simp add: Suc_diff_le) apply (metis EF LeftDerivation_implies_Derivation LeftDerivation_skip_prefix \' append_eq_conv_conj assms(3) drop_take is_word_Derivation_derivation_ge is_word_\ take_all take_append) using EF Suc_diff_le \' assms(3) j_ge_d by auto qed lemma LeftDerives1_propagate_prefix: "LeftDerives1 (\ @ \) i r \ \ i \ length \ \ is_prefix \ \" proof - assume a1: "LeftDerives1 (\ @ \) i r \" assume a2: "length \ \ i" have f3: "take i (\ @ \) = take i \" using a1 Derives1_take LeftDerives1_implies_Derives1 by blast then have f4: "length (take i \) = i" using a1 by (metis (no_types) Derives1_bound LeftDerives1_implies_Derives1 dual_order.strict_implies_order length_take min.absorb2) have "take (length \) (take i \) = \" using f3 a2 by (simp add: append_eq_conv_conj) then show ?thesis using f4 a2 by (metis (no_types) append_Nil2 append_eq_conv_conj diff_is_0_eq' is_prefix_take take_0 take_append) qed lemma LeftDerivationIntro_cut_prefix: assumes "LeftDerivationIntro (\@\) i r ix D j \" assumes "derivation_ge D (length \)" assumes "i \ length \" assumes is_word_\: "is_word \" shows "\ \'. \ = \ @ \' \ LeftDerivationIntro \ (i - length \) r ix (derivation_shift D (length \) 0) (j - length \) \'" proof - from iffD1[OF LeftDerivationIntro_def assms(1)] obtain \ where \: "LeftDerives1 (\ @ \) i r \ \ ix < length (snd r) \ snd r ! ix = \ ! j \ LeftDerivationFix \ (i + ix) D j \" by blast have "\ \'. \ = \ @ \'" using LeftDerives1_propagate_prefix \ assms(3) by (metis append_dropped_prefix) then obtain \' where \': "\ = \ @ \'" by blast with \ have "LeftDerives1 (\ @ \) i r (\ @ \')" by simp from LeftDerives1_skip_prefix[OF assms(3) this] have \_\': "LeftDerives1 \ (i - length \) r \'" by blast have ldfix: "LeftDerivationFix (\ @ \') (i + ix) D j \" using \ \' by auto have \_le_i_plus_ix: "length \ \ i + ix" using assms(3) by arith from LeftDerivationFix_cut_prefix[OF ldfix assms(2) \_le_i_plus_ix assms(4)] obtain \' where \': "\ = \ @ \' \ LeftDerivationFix \' (i + ix - length \) (derivation_shift D (length \) 0) (j - length \) \'" by blast have same_symbol: "\ ! j = \' ! (j - length \)" by (metis LeftDerivationFix_def \ \' \_le_i_plus_ix \' leD nth_append) have \'_\': "LeftDerivationFix \' (i - length \ + ix) (derivation_shift D (length \) 0) (j - length \) \'" by (simp add: \' assms(3)) show ?thesis apply (simp add: LeftDerivationIntro_def) apply (rule_tac x=\' in exI) apply (auto simp add: \') apply (rule_tac x=\' in exI) by (auto simp add: \ \_\' same_symbol \'_\') qed lemma LeftDerivationLadder_implies_LeftDerivation_at_index: assumes "LeftDerivationLadder \ D L \" assumes "index < length L" shows "LeftDerivation \ (take (ladder_n L index) D) (ladder_\ \ D L index)" using LeftDerivationLadder_def LeftDerivation_take_derive assms(1) ladder_\_def by auto lemma LeftDerivationLadder_cut_prefix_propagate: assumes ladder: "LeftDerivationLadder (\@\) D L \" assumes is_word_\: "is_word \" assumes derivation_ge_\: "derivation_ge D (length \)" assumes ladder_i_0: "ladder_i L 0 \ length \" assumes L': "L' = ladder_cut_prefix (length \) L" assumes D': "D' = derivation_shift D (length \) 0" shows "index < length L \ LeftDerivation \ (take (ladder_n L' index) D') (ladder_\ \ D' L' index) \ ladder_\ (\@\) D L index = \@(ladder_\ \ D' L' index) \ ladder_\ (\@\) D L index = \@(ladder_\ \ D' L' index)" proof (induct index) case 0 have ladder_\: "ladder_\ (\@\) D L 0 = \@(ladder_\ \ D' L' 0)" by (simp add: ladder_\_def) have ldfix: "LeftDerivationFix (\@\) (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ (\@\) D L 0)" using ladder LeftDerivationLadder_def by blast have dge_take: "derivation_ge (take (ladder_n L 0) D) (length \)" using derivation_ge_\ by (metis append_take_drop_id derivation_ge_append) from LeftDerivationFix_cut_prefix[OF ldfix dge_take ladder_i_0 is_word_\] obtain \' where \': "ladder_\ (\ @ \) D L 0 = \ @ \' \ LeftDerivationFix \ (ladder_i L 0 - length \) (derivation_shift (take (ladder_n L 0) D) (length \) 0) (ladder_j L 0 - length \) \'" by blast have ladder_\: "ladder_\ (\@\) D L 0 = \@(ladder_\ \ D' L' 0)" using \' by (metis "0.prems" D' Derive L' LeftDerivationFix_def LeftDerivation_implies_Derivation ladder_\_def ladder_cut_prefix_n take_derivation_shift) have "LeftDerivation \ (take (ladder_n L' 0) D') (ladder_\ \ D' L' 0)" proof - have "LeftDerivation (\@\) (take (ladder_n L 0) D) (ladder_\ (\@\) D L 0)" using LeftDerivationLadder_implies_LeftDerivation_at_index ladder "0.prems" by blast then show ?thesis by (metis D' LeftDerivationLadder_def LeftDerivation_skip_prefix LeftDerivation_take_derive derivation_ge_\ ladder ladder_\_def) qed then show ?case using ladder_\ ladder_\ by auto next case (Suc index) have index_less_L: "index < length L" using Suc(2) by arith then have induct: "ladder_\ (\@\) D L index = \@(ladder_\ \ D' L' index)" using Suc by blast then have ladder_\: "ladder_\ (\@\) D L (Suc index) = \@(ladder_\ \ D' L' (Suc index))" by (simp add: ladder_\_def) have introsAt: "LeftDerivationIntrosAt (\@\) D L (Suc index)" using Suc(2) ladder by (metis LeftDerivationIntros_def LeftDerivationLadder_def Suc_eq_plus1_left le_add1) obtain n m e E where n: "n = ladder_n L (Suc index - 1)" and m: "m = ladder_n L (Suc index)" and e: "e = D ! n" and E: "E = drop (Suc n) (take m D)" by blast from iffD1[OF LeftDerivationIntrosAt_def introsAt] have "LeftDerivationIntro (ladder_\ (\ @ \) D L (Suc index)) (ladder_i L (Suc index)) (snd e) (ladder_ix L (Suc index)) E (ladder_j L (Suc index)) (ladder_\ (\ @ \) D L (Suc index))" using n m e E Let_def by meson then have ldintro: "LeftDerivationIntro (\@(ladder_\ \ D' L' (Suc index))) (ladder_i L (Suc index)) (snd e) (ladder_ix L (Suc index)) E (ladder_j L (Suc index)) (ladder_\ (\ @ \) D L (Suc index))" by (simp add: ladder_\) have dge_E_\: "derivation_ge E (length \)" apply (simp add: E) using derivation_ge_\ by (metis append_take_drop_id derivation_ge_append) have ladder_i_Suc: "length \ \ ladder_i L (Suc index)" using Suc.prems derivation_ge_LeftDerivationLadder derivation_ge_\ ladder ladder_i_0 by blast from LeftDerivationIntro_cut_prefix[OF ldintro dge_E_\ ladder_i_Suc is_word_\] obtain \' where \': "ladder_\ (\ @ \) D L (Suc index) = \ @ \' \ LeftDerivationIntro (ladder_\ \ D' L' (Suc index)) (ladder_i L (Suc index) - length \) (snd e) (ladder_ix L (Suc index)) (derivation_shift E (length \) 0) (ladder_j L (Suc index) - length \) \'" by blast then have "LeftDerivation (ladder_\ \ D' L' (Suc index)) ((ladder_i L (Suc index) - length \, snd e) # (derivation_shift E (length \) 0)) \'" using LeftDerivationIntro_implies_LeftDerivation by blast then have "LeftDerivation (ladder_\ \ D' L' index) ((ladder_i L (Suc index) - length \, snd e) # (derivation_shift E (length \) 0)) \'" by (auto simp add: ladder_\_def) have ld: "LeftDerivation \ (take (ladder_n L' (Suc index)) D') (ladder_\ \ D' L' (Suc index))" proof - have "LeftDerivation (\@\) (take (ladder_n L (Suc index)) D) (ladder_\ (\@\) D L (Suc index))" using LeftDerivationLadder_implies_LeftDerivation_at_index ladder Suc.prems by blast then show ?thesis by (metis D' LeftDerivationLadder_def LeftDerivation_skip_prefix LeftDerivation_take_derive derivation_ge_\ ladder ladder_\_def) qed then show ?case using \' D' Derive L' LeftDerivationIntro_def n m e E ld LeftDerivation_implies_Derivation ladder_\_def ladder_cut_prefix_n take_derivation_shift by (metis (no_types, lifting) LeftDerivationLadder_implies_LeftDerivation_at_index LeftDerivation_skip_prefixword_ex Suc.prems Suc_leI index_less_L is_word_\ ladder ladder_\ le_0_eq neq0_conv) qed theorem LeftDerivationLadder_cut_prefix: assumes ladder: "LeftDerivationLadder (\@\) D L \" assumes is_word_\: "is_word \" assumes ladder_i_0: "ladder_i L 0 \ length \" shows "\ D' L' \'. \ = \ @ \' \ LeftDerivationLadder \ D' L' \' \ D' = derivation_shift D (length \) 0 \ length L' = length L \ ladder_i L' 0 + length \ = ladder_i L 0 \ ladder_last_j L' + length \ = ladder_last_j L" proof - obtain D' where D': "D' = derivation_shift D (length \) 0" by blast obtain L' where L': "L' = ladder_cut_prefix (length \) L" by blast obtain \' where \': "\' = drop (length \) \" by blast have ladder_last_j_upper_bound: "ladder_last_j L < length \" using ladder using ladder_last_j_bound by blast have derivation_ge_\: "derivation_ge D (length \)" using is_word_\ LeftDerivationLadder_def LeftDerivation_implies_Derivation is_word_Derivation_derivation_ge ladder by blast note derivation_ge_ladder = derivation_ge_LeftDerivationLadder[OF derivation_ge_\ ladder ladder_i_0] have ladder_last_j_lower_bound: "ladder_last_j L \ length \" using LeftDerivationLadder_def derivation_ge_ladder is_ladder_def ladder ladder_last_j_def by auto from ladder_last_j_upper_bound ladder_last_j_lower_bound have \_less_\: "length \ < length \" by arith then have \_def: "\ = \ @ \'" by (metis LeftDerivation.simps(1) LeftDerivationLadder_def LeftDerivation_ge_take \' append_eq_conv_conj derivation_ge_\ ladder) have length_L_nonzero: "length L \ 0" using LeftDerivationLadder_def is_ladder_def ladder by auto have ladder_i_L'_thm: "\ index. index < length L \ ladder_i L' index + length \ = ladder_i L index" apply (simp add: L') apply (subst ladder_cut_prefix_i) apply simp using length_L_nonzero apply blast using derivation_ge_ladder by auto have ladder_j_L'_thm: "\ index. index < length L \ ladder_j L' index + length \ = ladder_j L index" apply (simp add: L') apply (subst ladder_cut_prefix_j) using LeftDerivationLadder_def is_ladder_def ladder apply blast using LeftDerivationLadder_def is_ladder_def ladder apply blast using derivation_ge_ladder by auto have length_L': "length L' = length L" using L' ladder_cut_prefix_length by simp have \_\': "LeftDerivation \ D' \'" using D' LeftDerivationLadder_def LeftDerivation_skip_prefix \' derivation_ge_\ ladder by blast have length_D': "length D' = length D" by (simp add: D') have is_ladder_D_L: "is_ladder D L" using LeftDerivationLadder_def ladder by blast { fix u :: nat assume u_bound_L': "u < length L'" have u_bound_L: "u < length L" using length_L' using u_bound_L' by simp have "ladder_n L' u \ length D'" apply (simp add: length_D' L') apply (subst ladder_cut_prefix_n) apply (simp add: u_bound_L) using length_L_nonzero apply arith using u_bound_L is_ladder_D_L by (simp add: is_ladder_def) } note is_ladder_1 = this { fix u :: nat fix v :: nat assume u_less_v: "u < v" assume v_bound_L': "v < length L'" then have v_bound_L: "v < length L" by (simp add: length_L') with u_less_v have u_bound_L: "u < length L" by arith have "ladder_n L' u < ladder_n L' v" apply (simp add: L') apply (subst ladder_cut_prefix_n) using u_bound_L apply blast using length_L_nonzero apply blast apply (subst ladder_cut_prefix_n) using v_bound_L apply blast using length_L_nonzero apply blast using u_less_v v_bound_L is_ladder_D_L by (simp add: is_ladder_def) } note is_ladder_2 = this have is_ladder_3: "ladder_last_n L' = length D'" apply (simp add: length_D' ladder_last_n_def L') apply (subst ladder_cut_prefix_n) apply (simp add: ladder_cut_prefix_length) using length_L_nonzero apply auto[1] using length_L_nonzero apply blast apply (simp add: ladder_cut_prefix_length) using is_ladder_D_L by (simp add: is_ladder_def ladder_last_n_def) have is_ladder_4: "LeftDerivationFix \ (ladder_i L' 0) (take (ladder_n L' 0) D') (ladder_j L' 0) (ladder_\ \ D' L' 0)" proof - have ldfix: "LeftDerivationFix (\@\) (ladder_i L 0) (take (ladder_n L 0) D) (ladder_j L 0) (ladder_\ (\@\) D L 0)" using ladder LeftDerivationLadder_def by blast have dge: "derivation_ge (take (ladder_n L 0) D) (length \)" using derivation_ge_\ by (metis append_take_drop_id derivation_ge_append) from LeftDerivationFix_cut_prefix[OF ldfix dge ladder_i_0 is_word_\] obtain \' where \': "ladder_\ (\ @ \) D L 0 = \ @ \' \ LeftDerivationFix \ (ladder_i L 0 - length \) (derivation_shift (take (ladder_n L 0) D) (length \) 0) (ladder_j L 0 - length \) \'" by blast then show ?thesis using LeftDerivationLadder_cut_prefix_propagate D' L' append_eq_conv_conj derivation_ge_\ is_word_\ ladder ladder_cut_prefix_i ladder_cut_prefix_j ladder_cut_prefix_n ladder_i_0 length_0_conv length_L_nonzero length_greater_0_conv take_derivation_shift by auto qed { fix index :: nat assume index_lower_bound: "Suc 0 \ index" assume index_upper_bound: "index < length L'" have introsAt: "LeftDerivationIntrosAt (\@\) D L index" by (metis LeftDerivationIntros_def LeftDerivationLadder_def One_nat_def index_lower_bound index_upper_bound ladder length_L') then have ladder_i_L: "ladder_i L index = fst (D ! ladder_n L (index - Suc 0))" by (metis LeftDerivationIntrosAt_def One_nat_def \LeftDerivationIntrosAt (\ @ \) D L index\) have ladder_i_L'_as_L: "ladder_i L' index = ladder_i L index - (length \)" using ladder_cut_prefix_i L' index_upper_bound is_ladder_D_L is_ladder_not_empty length_L' length_greater_0_conv by auto have length_L_gr_0: "length L > 0" using length_L' length_L_nonzero by arith have index_Suc_upper_bound_L: "index - Suc 0 < length L" using index_upper_bound length_L' by arith have "fst (D' ! ladder_n L' (index - Suc 0)) = fst (D ! ladder_n L (index - Suc 0)) - (length \)" apply (subst D', subst L') apply (subst ladder_cut_prefix_n[OF index_Suc_upper_bound_L length_L_gr_0]) apply (simp add: derivation_shift_def) using index_lower_bound index_upper_bound is_ladder_D_L ladder_n_minus_1_bound length_L' by auto then have ladder_i_L': "ladder_i L' index = fst (D' ! ladder_n L' (index - Suc 0))" using ladder_i_L ladder_i_L'_as_L by auto have "LeftDerivationIntro (ladder_\ \ D' L' index) (ladder_i L' index) (snd (D' ! ladder_n L' (index - Suc 0))) (ladder_ix L' index) (drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) D')) (ladder_j L' index) (ladder_\ \ D' L' index)" proof - have "LeftDerivationIntro (ladder_\ (\@\) D L index) (ladder_i L index) (snd (D ! ladder_n L (index - Suc 0))) (ladder_ix L index) (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D)) (ladder_j L index) (ladder_\ (\@\) D L index)" using introsAt by (metis LeftDerivationIntrosAt_def One_nat_def) then have ldintro: "LeftDerivationIntro (\@(ladder_\ \ D' L' index)) (ladder_i L index) (snd (D ! ladder_n L (index - Suc 0))) (ladder_ix L index) (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D)) (ladder_j L index) (ladder_\ (\@\) D L index)" using D' L' LeftDerivationLadder_cut_prefix_propagate derivation_ge_\ index_upper_bound is_word_\ ladder ladder_i_0 length_L' by auto have dge: "derivation_ge (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D)) (length \)" using derivation_ge_\ by (metis append_take_drop_id derivation_ge_append) have \_le_i_L: "length \ \ ladder_i L index" using derivation_ge_ladder index_upper_bound length_L' by auto from LeftDerivationIntro_cut_prefix[OF ldintro dge \_le_i_L is_word_\] obtain \' where \': "ladder_\ (\ @ \) D L index = \ @ \' \ LeftDerivationIntro (ladder_\ \ D' L' index) (ladder_i L index - length \) (snd (D ! ladder_n L (index - Suc 0))) (ladder_ix L index) (derivation_shift (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D)) (length \) 0) (ladder_j L index - length \) \'" by blast have h1: "ladder_i L' index = ladder_i L index - length \" using L' ladder_cut_prefix_i ladder_i_L'_as_L by blast have h2: "(snd (D' ! ladder_n L' (index - Suc 0))) = (snd (D ! ladder_n L (index - Suc 0)))" apply (subst L', subst ladder_cut_prefix_n) apply (simp add: index_Suc_upper_bound_L) using length_L_gr_0 apply auto[1] apply (subst D') apply (simp add: derivation_shift_def) using index_lower_bound index_upper_bound is_ladder_D_L ladder_n_minus_1_bound length_L' by auto have h3: "ladder_ix L' index = ladder_ix L index" using ladder_cut_prefix_ix L' index_upper_bound length_L' length_L_gr_0 by auto have h4: "(drop (Suc (ladder_n L' (index - Suc 0))) (take (ladder_n L' index) D')) = (derivation_shift (drop (Suc (ladder_n L (index - Suc 0))) (take (ladder_n L index) D)) (length \) 0)" apply (subst D') apply (subst L')+ apply (subst ladder_cut_prefix_n, simp add: index_Suc_upper_bound_L) using length_L_gr_0 apply blast apply (subst ladder_cut_prefix_n) using index_upper_bound length_L' apply arith using length_L_gr_0 apply blast apply (simp add: derivation_shift_def) by (simp add: drop_map take_map) have h5: "ladder_j L' index = ladder_j L index - length \" using ladder_cut_prefix_j L' index_upper_bound length_L' length_L_gr_0 by auto have h6: "ladder_\ \ D' L' index = \'" using D' L' LeftDerivationLadder_cut_prefix_propagate \' derivation_ge_\ index_upper_bound is_word_\ ladder ladder_i_0 length_L' by auto show ?thesis using h1 h2 h3 h4 h5 h6 \' by simp qed then have "LeftDerivationIntrosAt \ D' L' index" apply (auto simp add: LeftDerivationIntrosAt_def Let_def) using ladder_i_L' by blast } note is_ladder_5 = this show ?thesis apply (rule_tac x="D'" in exI) apply (rule_tac x="L'" in exI) apply (rule_tac x="\'" in exI) apply auto using \_def apply blast defer 1 using D' apply blast using L' ladder_cut_prefix_length apply auto[1] apply (subst ladder_i_L'_thm) using LeftDerivationLadder_def is_ladder_def ladder apply blast apply simp apply (simp add: ladder_last_j_def) apply (subst ladder_j_L'_thm) apply (simp add: length_L') using length_L_nonzero apply arith apply (simp add: length_L') apply (auto simp add: LeftDerivationLadder_def) using \_\' apply blast apply (auto simp add: is_ladder_def) using length_L_nonzero length_L' apply auto[1] using is_ladder_1 apply blast using is_ladder_2 apply blast using is_ladder_3 apply blast using is_ladder_4 apply blast by (auto simp add: LeftDerivationIntros_def is_ladder_5) qed end end diff --git a/thys/LocalLexing/TheoremD10.thy b/thys/LocalLexing/TheoremD10.thy --- a/thys/LocalLexing/TheoremD10.thy +++ b/thys/LocalLexing/TheoremD10.thy @@ -1,540 +1,540 @@ theory TheoremD10 imports TheoremD9 Ladder begin context LocalLexing begin lemma \

_wellformed: "p \ \

k u \ wellformed_tokens p" using \

_are_admissible admissible_wellformed_tokens by blast lemma \_token_length: "t \ \ k \ k + length (chars_of_token t) \ length Doc" -by (metis Nat.le_diff_conv2 \_is_prefix add.commute chars_of_token_def empty_\ +by (metis le_diff_conv2 \_is_prefix add.commute chars_of_token_def empty_\ empty_iff is_prefix_length le_neq_implies_less length_drop linear) lemma mono_Scan: "mono (Scan T k)" by (simp add: Scan_regular regular_implies_mono) lemma \_apply_setmonotone: "x \ I \ x \ \ k T I" using Complete_subset_\ LocalLexing.Complete_def LocalLexing_axioms by blast lemma Scan_apply_setmonotone: "x \ I \ x \ Scan T k I" by (simp add: Scan_def) lemma leftderives_padfront: assumes "leftderives \ \" assumes "is_word u" shows "leftderives (u@\) (u@\)" using LeftDerivation_append_prefix LeftDerivation_implies_leftderives assms(1) assms(2) leftderives_implies_LeftDerivation by blast lemma leftderives_padback: assumes "leftderives \ \" assumes "is_sentence u" shows "leftderives (\@u) (\@u)" using LeftDerivation_append_suffix LeftDerivation_implies_leftderives assms(1) assms(2) leftderives_implies_LeftDerivation by blast lemma leftderives_pad: assumes \_\: "leftderives \ \" assumes is_word: "is_word u" assumes is_sentence: "is_sentence v" shows "leftderives (u@\@v) (u@\@v)" by (simp add: \_\ is_sentence is_word leftderives_padback leftderives_padfront) lemma leftderives_rule: assumes "(N, w) \ \" shows "leftderives [N] w" by (metis append_Nil append_Nil2 assms is_sentence_def is_word_terminals leftderives1_def leftderives1_implies_leftderives list.pred_inject(1) terminals_empty wellformed_tokens_empty_path) lemma leftderives_rule_step: assumes ld: "leftderives a (u@[N]@v)" assumes rule: "(N, w) \ \" assumes is_word: "is_word u" assumes is_sentence: "is_sentence v" shows "leftderives a (u@w@v)" proof - have N_w: "leftderives [N] w" using rule leftderives_rule by blast then have "leftderives (u@[N]@v) (u@w@v)" using leftderives_pad is_word is_sentence by blast then show "leftderives a (u@w@v)" using leftderives_trans ld by blast qed lemma leftderives_trans_step: assumes ld: "leftderives a (u@b@v)" assumes rule: "leftderives b c" assumes is_word: "is_word u" assumes is_sentence: "is_sentence v" shows "leftderives a (u@c@v)" proof - have "leftderives (u@b@v) (u@c@v)" using leftderives_pad ld rule is_word is_sentence by blast then show ?thesis using leftderives_trans ld by blast qed lemma charslength_of_prefix: assumes "is_prefix a b" shows "charslength a \ charslength b" by (simp add: assms is_prefix_chars is_prefix_length) lemma item_rhs_simp[simp]: "item_rhs (Item (N, \) d i j) = \" by (simp add: item_rhs_def) definition Prefixes :: "'a list \ 'a list set" where "Prefixes p = { q . is_prefix q p }" lemma \_wellformed: "p \ \ \ wellformed_tokens p" by (simp add: \_are_admissible admissible_wellformed_tokens) lemma Prefixes_reflexive[simp]: "p \ Prefixes p" by (simp add: Prefixes_def is_prefix_def) lemma Prefixes_is_prefix: "q \ Prefixes p = is_prefix q p" by (simp add: Prefixes_def) lemma prefixes_are_paths': "p \ \ \ is_prefix q p \ q \ \" using \

.simps(3) \_def prefixes_are_paths by blast lemma thmD10_ladder: "p \ \ \ charslength p = k \ X \ T \ T \ \ k \ (N, \@\) \ \ \ r \ length p \ leftderives [\] ((terminals (take r p))@[N]@\) \ LeftDerivationLadder \ D L (terminals ((drop r p)@[X])) \ ladder_last_j L = length (drop r p) \ k' = k + length (chars_of_token X) \ x = Item (N, \@\) (length \) (charslength (take r p)) k' \ I = items_le k' (\ k' {} (Scan T k (Gen (Prefixes p)))) \ x \ I" proof (induct "length L" arbitrary: N \ \ r \ D L x rule: less_induct) case less have item_origin_x_def: "item_origin x = (charslength (take r p))" by (simp add: less.prems(11)) then have x_k: "item_origin x \ k" using charslength.simps is_prefix_chars is_prefix_length is_prefix_take less.prems(2) by force have item_end_x_def: "item_end x = k'" by (simp add: less.prems(11)) have item_dot_x_def: "item_dot x = length \" by (simp add: less.prems(11)) have k'_upperbound: "k' \ length Doc" using \_token_length less.prems(10) less.prems(3) less.prems(4) by blast note item_def = less.prems(11) note k' = less.prems(10) note rule_dom = less.prems(5) note p_charslength = less.prems(2) note p_dom = less.prems(1) note r = less.prems(6) note leftderives_start = less.prems(7) note X_dom = less.prems(3) have wellformed_x: "wellformed_item x" apply (auto simp add: wellformed_item_def item_def rule_dom p_charslength) apply (subst k') apply (subst charslength.simps[symmetric]) apply (subst p_charslength[symmetric]) using item_origin_x_def p_charslength x_k apply linarith apply (rule k'_upperbound) done have leftderives_\: "leftderives \ (terminals ((drop r p)@[X]))" using LeftDerivationLadder_def LeftDerivation_implies_leftderives less.prems(8) by blast have is_sentence_drop_pX: "is_sentence (drop r (terminals p) @ [terminal_of_token X])" by (metis derives_is_sentence is_sentence_concat leftderives_\ leftderives_implies_derives rule_\_type rule_dom terminals_append terminals_drop terminals_singleton) have snd_item_rule_x: "snd (item_rule x) = \@\" by (simp add: item_def) from less have "is_ladder D L" using LeftDerivationLadder_def by blast then have "length L \ 0" by (simp add: is_ladder_not_empty) then have "length L = 1 \ length L > 1" by arith then show ?case proof (induct rule: disjCases2) case 1 have "\ i. LeftDerivationFix \ i D (length (drop r p)) (terminals ((drop r p)@[X]))" using "1.hyps" LeftDerivationLadder_L_0 less.prems(8) less.prems(9) by fastforce then obtain i where LDF: "LeftDerivationFix \ i D (length (drop r p)) (terminals ((drop r p)@[X]))" by blast from LeftDerivationFix_splits_at_derives[OF this] obtain U a1 a2 b1 b2 where decompose: "splits_at \ i a1 U a2 \ splits_at (terminals (drop r p @ [X])) (length (drop r p)) b1 U b2 \ derives a1 b1 \ derives a2 b2" by blast then have b1: "b1 = terminals (drop r p)" by (simp add: append_eq_conv_conj splits_at_def) with decompose have U: "U = fst X" by (metis length_terminals nth_append_length splits_at_def terminal_of_token_def terminals_append terminals_singleton) from decompose b1 U have b2: "b2 = []" by (simp add: splits_at_combine splits_at_def) have D: "LeftDerivation \ D (terminals ((drop r p)@[X]))" using LDF LeftDerivationLadder_def less.prems(8) by blast let ?y = "Item (item_rule x) (length a1) (item_origin x) k" have wellformed_y: "wellformed_item ?y" using wellformed_x apply (auto simp add: wellformed_item_def x_k) using k' k'_upperbound apply arith apply (simp add: item_rhs_def snd_item_rule_x) using decompose splits_at_def by (simp add: is_prefix_length trans_le_add1) have y_nonterminal: "item_nonterminal ?y = N" by (simp add: item_def item_nonterminal_def) have item_\_y: "item_\ ?y = a1" by (metis append_assoc append_eq_conv_conj append_take_drop_id decompose item.sel(1) item.sel(2) item_\_def item_rhs_def snd_item_rule_x splits_at_def) have pvalid_y: "pvalid p ?y" apply (auto simp add: pvalid_def) using p_dom \_wellformed apply blast using wellformed_y apply auto[1] apply (rule_tac x=r in exI) apply (auto simp add: r) using p_charslength apply simp using item_def apply simp apply (rule_tac x=\ in exI) using y_nonterminal apply simp using is_derivation_def leftderives_start apply auto[1] apply (simp add: item_\_y) using b1 decompose by auto let ?z = "inc_item ?y k'" have item_rhs_y: "item_rhs ?y = \@\" by (simp add: item_def item_rhs_def) have split_\: "\ = a1@[U]@a2" using decompose splits_at_combine by blast have next_symbol_y: "next_symbol ?y = Some(fst X)" by (auto simp add: next_symbol_def is_complete_def item_rhs_y split_\ U) have z_in_Scan_y: "?z \ Scan T k {?y}" apply (simp add: Scan_def) apply (rule disjI2) apply (rule_tac x="?y" in exI) apply (rule_tac x="fst X" in exI) apply (rule_tac x="snd X" in exI) apply (auto simp add: bin_def X_dom) using k' chars_of_token_def apply simp using next_symbol_y apply simp done from pvalid_y have "?y \ Gen(Prefixes p)" apply (simp add: Gen_def) apply (rule_tac x=p in exI) by (auto simp add: paths_le_def p_dom) then have "Scan T k {?y} \ Scan T k (Gen(Prefixes p))" apply (rule_tac monoD[OF mono_Scan]) apply blast done with z_in_Scan_y have z_in_Scan_Gen: "?z \ Scan T k (Gen(Prefixes p))" using rev_subsetD by blast have wellformed_z: "wellformed_item ?z" using k' k'_upperbound next_symbol_y wellformed_inc_item wellformed_y by auto have item_\_z: "item_\ ?z = a2@\" apply (simp add: item_\_def) apply (simp add: item_rhs_y split_\) done have item_end_z: "item_end ?z = k'" by simp have x_via_z: "x = inc_dot (length a2) ?z" by (simp add: inc_dot_def less.prems(11) split_\) have x_in_z: "x \ \ k' {} {?z}" apply (subst x_via_z) apply (rule_tac thmD6[OF wellformed_z item_\_z item_end_z]) using decompose b2 by blast have "\ k' {} {?z} \ \ k' {} (Scan T k (Gen(Prefixes p)))" apply (rule_tac monoD[OF mono_\]) using z_in_Scan_Gen using empty_subsetI insert_subset by blast then have x_in_Scan_Gen: "x \ \ k' {} (Scan T k (Gen(Prefixes p)))" using x_in_z by blast have "item_end x = k'" by (simp add: item_end_x_def) with x_in_Scan_Gen show ?case using items_le_def less.prems(12) mem_Collect_eq nat_le_linear by blast next case 2 obtain i where i: "i = ladder_i L 0" by blast obtain i' where i': "i' = ladder_j L 0" by blast obtain \' where \': "\' = ladder_\ \ D L 0" by blast obtain n where n: "n = ladder_n L 0" by blast have ldfix: "LeftDerivationFix \ i (take n D) i' \'" using LeftDerivationLadder_def \' i i' n less.prems(8) by blast have \'_alt: "\' = ladder_\ \ D L 1" using 2 by (simp add: \' ladder_\_def) have i'_alt: "i' = ladder_i L 1" using 2 by (simp add: i' ladder_i_def) obtain e where e: "e = D ! n" by blast obtain ix where ix: "ix = ladder_ix L 1" by blast obtain m where m: "m = ladder_n L 1" by blast obtain E where E: "E = drop (Suc n) (take m D)" by blast have ldintro: "LeftDerivationIntro \' i' (snd e) ix E (ladder_j L 1) (ladder_\ \ D L 1)" by (metis "2.hyps" LeftDerivationIntrosAt_def LeftDerivationIntros_def LeftDerivationLadder_def One_nat_def \'_alt E diff_Suc_1 e i'_alt ix leI less.prems(8) m n not_less_eq zero_less_one) have is_nonterminal_\'_at_i': "is_nonterminal (\' ! i')" using LeftDerivationIntro_implies_nonterminal ldintro by blast then have is_nonterminal_\_at_i: "is_nonterminal (\ ! i)" using LeftDerivationFix_def ldfix by auto then have "\ A a1 a2 a1'. splits_at \ i a1 A a2 \ splits_at \' i' a1' A a2 \ LeftDerivation a1 (take n D) a1'" using LeftDerivationFix_splits_at_nonterminal ldfix by auto then obtain A a1 a2 a1' where A: "splits_at \ i a1 A a2 \ splits_at \' i' a1' A a2 \ LeftDerivation a1 (take n D) a1'" by blast have A_def: "A = \' ! i'" using A splits_at_def by blast have is_nonterminal_A: "is_nonterminal A" using A_def is_nonterminal_\'_at_i' by blast have "\ rule. e = (i', rule)" by (metis e "2.hyps" LeftDerivationIntrosAt_def LeftDerivationIntros_def LeftDerivationLadder_def One_nat_def Suc_leI diff_Suc_1 i'_alt less.prems(8) n prod.collapse zero_less_one) then obtain rule where rule: "e = (i', rule)" by blast obtain w where w: "w = snd rule" by blast obtain \'' where \'': "\'' = a1' @ w @ a2" by blast have \'_\'': "LeftDerives1 \' i' rule \''" by (metis A w LeftDerivationFix_is_sentence LeftDerivationIntro_def LeftDerivationIntro_examine_rule LeftDerives1_def \'' ldfix ldintro prod.collapse rule snd_conv splits_at_implies_Derives1) then have is_word_a1': "is_word a1'" using LeftDerives1_splits_at_is_word A by blast have A_eq_fst_rule: "A = fst rule" using A LeftDerivationIntro_examine_rule ldintro rule by fastforce have ix_bound: "ix < length w" using ix w rule ldintro LeftDerivationIntro_def snd_conv by auto then have "\ w1 W w2. splits_at w ix w1 W w2" using splits_at_def by blast then obtain w1 W w2 where W: "splits_at w ix w1 W w2" by blast have a1'_w_a2: "a1'@w@a2 = ladder_stepdown_\_0 \ D L" using ladder_stepdown_\_0_altdef "2.hyps" A \'_alt e i'_alt less.prems(8) n rule snd_conv w by force from LeftDerivationLadder_stepdown[OF less.prems(8) 2] obtain L' where L': "LeftDerivationLadder (a1'@(w@a2)) (drop (ladder_stepdown_diff L) D) L' (terminals (drop r p @ [X])) \ length L' = length L - 1 \ ladder_i L' 0 = ladder_i L 1 + ladder_ix L 1 \ ladder_last_j L' = ladder_last_j L" using a1'_w_a2 by auto have ladder_i_L'_0: "ladder_i L' 0 = i' + ix" using L' i'_alt ix by auto have ladder_last_j_L': "ladder_last_j L' = length (drop r p)" using L' less.prems by auto from L' have this1: "LeftDerivationLadder (a1'@(w@a2)) (drop (ladder_stepdown_diff L) D) L' (terminals (drop r p @ [X]))" by blast have this2: "length a1' \ ladder_i L' 0" using A ladder_i_L'_0 splits_at_def by auto from LeftDerivationLadder_cut_prefix[OF this1 is_word_a1' this2] obtain D' L'' \' where L'': "terminals (drop r p @ [X]) = a1' @ \' \ LeftDerivationLadder (w @ a2) D' L'' \' \ D' = derivation_shift (drop (ladder_stepdown_diff L) D) (length a1') 0 \ length L'' = length L' \ ladder_i L'' 0 + length a1' = ladder_i L' 0 \ ladder_last_j L'' + length a1' = ladder_last_j L'" by blast have length_a1'_bound: "length a1' \ length (drop r p)" using L'' ladder_last_j_L' by linarith then have is_prefix_a1'_drop_r_p: "is_prefix a1' (terminals (drop r p))" proof - have f1: "\ss ssa ssb. \ is_prefix (ss::symbol list) (ssa @ ssb) \ is_prefix ss ssa \ (\ssc. ssc \ [] \ is_prefix ssc ssb \ ss = ssa @ ssc)" by (simp add: is_prefix_of_append) have f2: "\ss ssa. is_prefix ((ss::symbol list) @ ssa) ss \ \ is_prefix ssa []" by (metis (no_types) append_Nil2 is_prefix_cancel) have f3: "\ss. is_prefix ss [] \ \ is_prefix (terminals (drop r p) @ ss) a1'" by (metis (no_types) drop_eq_Nil is_prefix_append length_a1'_bound length_terminals) have "is_prefix a1' (a1' @ \') \ is_prefix a1' a1'" by (metis (no_types) append_Nil2 is_prefix_cancel is_prefix_empty) then show ?thesis using f3 f2 f1 by (metis L'' terminals_append) qed obtain r' where r': "r' = r + i'" by blast have length_a1'_eq_i': "length a1' = i'" using A less_or_eq_imp_le min.absorb2 splits_at_def by auto have a1'_r': "r \ r' \ r' \ length p \ terminals (drop r p) = a1' @ (terminals (drop r' p))" using is_prefix_a1'_drop_r_p r' proof - have "\ q. terminals (drop r p) = a1' @ q" using is_prefix_a1'_drop_r_p by (metis is_prefix_unsplit) then obtain q where q: "terminals (drop r p) = a1' @ q" by blast have "q = drop i' (terminals (drop r p))" using length_a1'_eq_i' q by (simp add: append_eq_conv_conj) then have "q = terminals (drop i' (drop r p))" by simp then have "q = terminals (drop r' p)" by (simp add: r' add.commute) with q show ?thesis using add.commute diff_add_inverse le_Suc_ex le_add1 le_diff_conv length_a1'_bound length_a1'_eq_i' length_drop r r' by auto qed have ladder_i_L'': "ladder_i L'' 0 = ix" using L'' ladder_i_L'_0 length_a1'_eq_i' add.commute add_left_cancel by linarith have L''_ladder: "LeftDerivationLadder (w @ a2) D' L'' \'" using L'' by blast have "ladder_i L'' 0 < length w" using ladder_i_L'' ix_bound by blast from LeftDerivationLadder_cut_appendix[OF L''_ladder this] obtain E' F' \1' \2' L''' where L''': "D' = E' @ F' \ \' = \1' @ \2' \ LeftDerivationLadder w E' L''' \1' \ derivation_ge F' (length \1') \ LeftDerivation a2 (derivation_shift F' (length \1') 0) \2' \ length L''' = length L'' \ ladder_i L''' 0 = ladder_i L'' 0 \ ladder_last_j L''' = ladder_last_j L''" by blast obtain z where z: "z = Item (A, w) (length w) (charslength (take r' p)) k'" by blast have z1: "length L''' < length L" using "2.hyps" L' L'' L''' by linarith have z2: "p \ \" by (simp add: p_dom) have z3: "charslength p = k" using p_charslength by auto have z4: "X \ T" by (simp add: X_dom) have z5: "T \ \ k" by (simp add: less.prems(4)) have "rule \ \" using Derives1_rule LeftDerives1_implies_Derives1 \'_\'' by blast then have z6: "(A, w @ []) \ \" using w A_eq_fst_rule by auto have z7: "r' \ length p" using a1'_r' by linarith have "leftderives [\] ((terminals (take r p))@[N]@\)" using leftderives_start by blast then have "leftderives [\] ((terminals (take r p))@(\@\)@\)" by (metis \_wellformed is_derivation_def is_derivation_is_sentence is_sentence_concat is_word_terminals_take leftderives_implies_derives leftderives_rule_step p_dom rule_dom) then have "leftderives [\] ((terminals (take r p))@a1@([A]@a2@\)@\)" using A splits_at_combine append_assoc by fastforce then have z8_helper: "leftderives [\] ((terminals (take r p))@a1'@([A]@a2@\)@\)" by (meson A LeftDerivation_implies_leftderives \_wellformed is_derivation_def is_derivation_is_sentence is_sentence_concat is_word_terminals_take leftderives_implies_derives leftderives_trans_step p_dom) have join_terminals: "(terminals (take r p))@a1' = terminals (take r' p)" by (metis is_prefix_a1'_drop_r_p length_a1'_eq_i' r' take_add take_prefix terminals_drop terminals_take) from z8_helper join_terminals have z8: "leftderives [\] (terminals (take r' p) @ [A] @ a2 @ \ @ \)" by (metis append_assoc) have \'_altdef: "\' = terminals (drop r' p @ [X])" using L'' a1'_r' by auto have "ladder_last_j L''' + length a1' = length (drop r p)" using L'' L''' ladder_last_j_L' by linarith then have ladder_last_j_L'''_\': "ladder_last_j L''' = length \' - 1" by (simp add: \'_altdef length_a1'_eq_i' r') then have "length \' - 1 < length \1'" using L''' ladder_last_j_bound by fastforce then have "length \1' + length \2' - 1 < length \1'" using L''' by simp then have "length \2' = 0" by arith then have \2': "\2' = []" by simp then have \1': "\1' = terminals (drop r' p @ [X])" using \'_altdef L''' by auto then have z9: "LeftDerivationLadder w E' L''' (terminals (drop r' p @ [X]))" using L''' by blast have z10: "ladder_last_j L''' = length (drop r' p)" using \'_altdef ladder_last_j_L'''_\' by auto note z11 = k' have z12: "z = Item (A, w @ []) (length w) (charslength (take r' p)) k'" using z by simp note z13 = less.prems(12) note induct = less.hyps(1)[of L''' A w "[]" r' "a2@\@\" E' z] note z_in_I = induct[OF z1 z2 z3 z4 z5 z6 z7 z8 z9 z10 z11 z12 z13] have a2_derives_empty: "derives a2 []" using L''' \2' using LeftDerivation_implies_leftderives leftderives_implies_derives by blast obtain x1 where x1: "x1 = Item (N, \@\) (length a1) (charslength (take r p)) (charslength (take r' p))" by blast obtain q where q: "q = take r' p" by blast then have is_prefix_q_p: "is_prefix q p" by simp then have q_in_Prefixes: "q \ Prefixes p" using Prefixes_is_prefix by blast then have wellformed_q: "wellformed_tokens q" using p_dom is_prefix_q_p prefixes_are_paths' \_wellformed by blast have item_rule_x1: "item_rule x1 = (N, \@\)" using x1 by simp have is_prefix_r_r': "is_prefix (take r p) (take r' p)" by (metis append_eq_conv_conj is_prefix_take r' take_add) then have charslength_le_r_r': "charslength (take r p) \ charslength (take r' p)" using charslength_of_prefix by blast have "is_prefix (take r' p) p" by auto then have charslength_r'_p: "charslength (take r' p) \ charslength p" using charslength_of_prefix by blast have "charslength p \ length Doc" using less.prems(1) add_leE k' k'_upperbound z3 by blast with charslength_r'_p have charslength_r'_Doc: "charslength (take r' p) \ length Doc" by arith have \_decompose: "\ = a1 @ [A] @ a2" using A splits_at_combine by blast have wellformed_x1: "wellformed_item x1" apply (auto simp add: wellformed_item_def) using item_rule_x1 less.prems apply auto[1] using charslength_le_r_r' x1 apply auto[1] using charslength_r'_Doc x1 apply auto[1] using x1 \_decompose by simp have item_nonterminal_x1: "item_nonterminal x1 = N" by (simp add: x1 item_nonterminal_def) have r_q_p: "take r (terminals q) = terminals (take r p)" by (metis q is_prefix_r_r' length_take min.absorb2 r take_prefix terminals_take) have item_\_x1: "item_\ x1 = a1" by (simp add: \_decompose item_\_def x1) have a1': "a1' = drop r (terminals q)" by (metis append_eq_conv_conj join_terminals length_take length_terminals min.absorb2 q r) have pvalid_q_x1: "pvalid q x1" apply (simp add: pvalid_def wellformed_q wellformed_x1 item_nonterminal_x1) apply (rule_tac x=r in exI) apply auto apply (simp add: a1'_r' min.absorb2 q) apply (simp add: q x1) apply (simp add: q x1 r') using r_q_p less.prems(7) append_Cons is_leftderivation_def leftderivation_implies_derivation apply fastforce apply (simp add: item_\_x1) using a1' A LeftDerivation_implies_leftderives leftderives_implies_derives by blast have item_end_x1_le_k': "item_end x1 \ k'" using charslength_r'_p apply (simp add: x1) using less.prems by auto have x1_in_I: "x1 \ I" apply (subst less.prems(12)) apply (auto simp add: items_le_def item_end_x1_le_k') apply (rule \_apply_setmonotone) apply (rule Scan_apply_setmonotone) apply (simp add: Gen_def) apply (rule_tac x=q in exI) by (simp add: pvalid_q_x1 paths_le_def q_in_Prefixes) obtain x2 where x2: "x2 = inc_item x1 k'" by blast have x1_in_bin: "x1 \ bin I (item_origin z)" using bin_def x1 x1_in_I z12 by auto have x2_in_Complete: "x2 \ Complete k' I" apply (simp add: Complete_def) apply (rule disjI2) apply (rule_tac x=x1 in exI) apply (simp add: x2) apply (rule_tac x=z in exI) apply auto using x1_in_bin apply blast using bin_def z12 z_in_I apply auto[1] apply (simp add: is_complete_def z12) by (simp add: \_decompose is_complete_def item_nonterminal_def next_symbol_def x1 z12) have wf_I': "wellformed_items (\ k' {} (Scan T k (Gen (Prefixes p))))" by (simp add: wellformed_items_Gen wellformed_items_Scan wellformed_items_\ z5) from items_le_Complete[OF this] x2_in_Complete have x2_in_I: "x2 \ I" by (metis Complete_\_fix z13) have "wellformed_items I" using wf_I' items_le_is_filter wellformed_items_def z13 by auto with x2_in_I have wellformed_x2: "wellformed_item x2" by (simp add: wellformed_items_def) have item_dot_x2: "item_dot x2 = Suc (length a1)" by (simp add: x2 x1) have item_\_x2: "item_\ x2 = a2 @ \" apply (simp add: item_\_def item_dot_x2) apply (simp add: item_rhs_def x2 x1 \_decompose) done have item_end_x2: "item_end x2 = k'" by (simp add: x2) note inc_dot_x2_by_a2 = thmD6[OF wellformed_x2 item_\_x2 item_end_x2 a2_derives_empty] have "x = inc_dot (length a2) x2" by (simp add: \_decompose inc_dot_def less.prems(11) x1 x2) with inc_dot_x2_by_a2 have "x \ \ k' {} {x2}" by auto then have "x \ \ k' {} I" using x2_in_I by (meson contra_subsetD empty_subsetI insert_subset monoD mono_\) then show "x \ I" by (metis (no_types, lifting) \_subset_elem_trans dual_order.refl item_end_x_def items_le_def items_le_is_filter mem_Collect_eq z13) qed qed theorem thmD10: assumes p_dom: "p \ \" assumes p_charslength: "charslength p = k" assumes X_dom: "X \ T" assumes T_dom: "T \ \ k" assumes rule_dom: "(N, \@\) \ \" assumes r: "r \ length p" assumes leftderives_start: "leftderives [\] ((terminals (take r p))@[N]@\)" assumes leftderives_\: "leftderives \ (terminals ((drop r p)@[X]))" assumes k': "k' = k + length (chars_of_token X)" assumes item_def: "x = Item (N, \@\) (length \) (charslength (take r p)) k'" assumes I: "I = items_le k' (\ k' {} (Scan T k (Gen (Prefixes p))))" shows "x \ I" proof - have "\ D. LeftDerivation \ D (terminals ((drop r p)@[X]))" using leftderives_\ leftderives_implies_LeftDerivation by blast then obtain D where D: "LeftDerivation \ D (terminals ((drop r p)@[X]))" by blast have is_sentence: "is_sentence (terminals (drop r p @ [X]))" using derives_is_sentence is_sentence_concat leftderives_\ leftderives_implies_derives rule_\_type rule_dom by blast have "\ L. LeftDerivationLadder \ D L (terminals ((drop r p)@[X])) \ ladder_last_j L = length (drop r p)" apply (rule LeftDerivationLadder_exists) apply (rule D) apply (rule is_sentence) by auto then obtain L where L: "LeftDerivationLadder \ D L (terminals ((drop r p)@[X]))" and L_ladder_last_j: "ladder_last_j L = length (drop r p)" by blast from thmD10_ladder[OF assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) assms(7) L L_ladder_last_j k' item_def I] show ?thesis . qed end end diff --git a/thys/Probabilistic_Timed_Automata/PTA_Reachability.thy b/thys/Probabilistic_Timed_Automata/PTA_Reachability.thy --- a/thys/Probabilistic_Timed_Automata/PTA_Reachability.thy +++ b/thys/Probabilistic_Timed_Automata/PTA_Reachability.thy @@ -1,2268 +1,2268 @@ theory PTA_Reachability imports PTA begin section \Classifying Regions for Divergence\ subsection \Pairwise\ coinductive pairwise :: "('a \ 'a \ bool) \ 'a stream \ bool" for P where "P a b \ pairwise P (b ## xs) \ pairwise P (a ## b ## xs)" lemma pairwise_Suc: "pairwise P xs \ P (xs !! i) (xs !! (Suc i))" by (induction i arbitrary: xs) (force elim: pairwise.cases)+ lemma Suc_pairwise: "\ i. P (xs !! i) (xs !! (Suc i)) \ pairwise P xs" apply (coinduction arbitrary: xs) apply (subst stream.collapse[symmetric]) apply (rewrite in "stl _" stream.collapse[symmetric]) apply (intro exI conjI, rule HOL.refl) apply (erule allE[where x = 0]; simp; fail) by simp (metis snth.simps(2)) lemma pairwise_iff: "pairwise P xs \ (\ i. P (xs !! i) (xs !! (Suc i)))" using pairwise_Suc Suc_pairwise by blast lemma pairwise_stlD: "pairwise P xs \ pairwise P (stl xs)" by (auto elim: pairwise.cases) lemma pairwise_pairD: "pairwise P xs \ P (shd xs) (shd (stl xs))" by (auto elim: pairwise.cases) lemma pairwise_mp: assumes "pairwise P xs" and lift: "\ x y. x \ sset xs \ y \ sset xs \ P x y \ Q x y" shows "pairwise Q xs" using assms apply (coinduction arbitrary: xs) subgoal for xs apply (subst stream.collapse[symmetric]) apply (rewrite in "stl _" stream.collapse[symmetric]) apply (intro exI conjI) apply (rule HOL.refl) by (auto intro: stl_sset dest: pairwise_pairD pairwise_stlD) done lemma pairwise_sdropD: "pairwise P (sdrop i xs)" if "pairwise P xs" using that proof (coinduction arbitrary: i xs) case (pairwise i xs) then show ?case apply (inst_existentials "shd (sdrop i xs)" "shd (stl (sdrop i xs))" "stl (stl (sdrop i xs))") subgoal by (auto dest: pairwise_Suc) (metis sdrop_simps(1) sdrop_stl stream.collapse) subgoal by (inst_existentials "i - 1" "stl xs") (auto dest: pairwise_Suc pairwise_stlD) by (metis sdrop_simps(2) stream.collapse) qed subsection \Regions\ (* XXX Move. Rename? *) lemma gt_GreaterD: assumes "u \ region X I r" "valid_region X k I r" "c \ X" "u c > k c" shows "I c = Greater (k c)" proof - from assms have "intv_elem c u (I c)" "valid_intv (k c) (I c)" by auto with assms(4) show ?thesis by (cases "I c") auto qed lemma const_ConstD: assumes "u \ region X I r" "valid_region X k I r" "c \ X" "u c = d" "d \ k c" shows "I c = Const d" proof - from assms have "intv_elem c u (I c)" "valid_intv (k c) (I c)" by auto with assms(4,5) show ?thesis by (cases "I c") auto qed lemma not_Greater_bounded: assumes "I x \ Greater (k x)" "x \ X" "valid_region X k I r" "u \ region X I r" shows "u x \ k x" proof - from assms have "intv_elem x u (I x)" "valid_intv (k x) (I x)" by auto with assms(1) show "u x \ k x" by (cases "I x") auto qed lemma Greater_closed: fixes t :: "real" assumes "u \ region X I r" "valid_region X k I r" "c \ X" "I c = Greater (k c)" "t > k c" shows "u(c := t) \ region X I r" using assms apply (intro region.intros) apply (auto; fail) apply standard subgoal for x by (cases "x = c"; cases "I x"; force intro!: intv_elem.intros) by auto lemma Greater_unbounded_aux: assumes "finite X" "valid_region X k I r" "c \ X" "I c = Greater (k c)" shows "\ u \ region X I r. u c > t" using assms Greater_closed[OF _ assms(2-4)] proof - let ?R = "region X I r" let ?t = "if t > k c then t + 1 else k c + 1" have t: "?t > k c" by auto from region_not_empty[OF assms(1,2)] obtain u where u: "u \ ?R" by auto from Greater_closed[OF this assms(2-4) t] have "u(c:=?t) \ ?R" by auto with t show ?thesis by (inst_existentials "u(c:=?t)") auto qed subsection \Unbounded and Zero Regions\ definition "unbounded x R \ \ t. \ u \ R. u x > t" definition "zero x R \ \ u \ R. u x = 0" lemma Greater_unbounded: assumes "finite X" "valid_region X k I r" "c \ X" "I c = Greater (k c)" shows "unbounded c (region X I r)" using Greater_unbounded_aux[OF assms] unfolding unbounded_def by blast lemma unbounded_Greater: assumes "valid_region X k I r" "c \ X" "unbounded c (region X I r)" shows "I c = Greater (k c)" using assms unfolding unbounded_def by (auto intro: gt_GreaterD) lemma Const_zero: assumes "c \ X" "I c = Const 0" shows "zero c (region X I r)" using assms unfolding zero_def by force lemma zero_Const: assumes "finite X" "valid_region X k I r" "c \ X" "zero c (region X I r)" shows "I c = Const 0" proof - from assms obtain u where "u \ region X I r" by atomize_elim (auto intro: region_not_empty) with assms show ?thesis unfolding zero_def by (auto intro: const_ConstD) qed lemma zero_all: assumes "finite X" "valid_region X k I r" "c \ X" "u \ region X I r" "u c = 0" shows "zero c (region X I r)" proof - from assms have "intv_elem c u (I c)" "valid_intv (k c) (I c)" by auto then have "I c = Const 0" using assms(5) by cases auto with assms have "u' c = 0" if "u' \ region X I r" for u' using that by force then show ?thesis unfolding zero_def by blast qed section \Reachability\ subsection \Definitions\ locale Probabilistic_Timed_Automaton_Regions_Reachability = Probabilistic_Timed_Automaton_Regions k v n not_in_X A for k v n not_in_X and A :: "('c, t, 's) pta" + fixes \ \ :: "('s * ('c, t) cval) \ bool" fixes s assumes \: "\ x y. x \ S \ x ~ y \ \ x \ \ y" assumes \: "\ x y. x \ S \ x ~ y \ \ x \ \ y" assumes s[intro, simp]: "s \ S" begin definition "\' \ absp \" definition "\' \ absp \" definition "s' \ abss s" lemma s_s'_cfg_on[intro]: assumes "cfg \ MDP.cfg_on s" shows "absc cfg \ R_G.cfg_on s'" proof - from assms s have "cfg \ valid_cfg" unfolding MDP.valid_cfg_def by auto then have "absc cfg \ R_G.cfg_on (state (absc cfg))" by (auto intro: R_G.valid_cfgD) with assms show ?thesis unfolding s'_def by (auto simp: state_absc) qed lemma s'_\[simp, intro]: "s' \ \" unfolding s'_def using s by auto lemma s'_s_cfg_on[intro]: assumes "cfg \ R_G.cfg_on s'" shows "repcs s cfg \ MDP.cfg_on s" proof - from assms s have "cfg \ R_G.valid_cfg" unfolding R_G.valid_cfg_def by auto with assms have "repcs s cfg \ valid_cfg" by (auto simp: s'_def intro: R_G.valid_cfgD) then show ?thesis by (auto dest: MDP.valid_cfgD) qed lemma (in Probabilistic_Timed_Automaton_Regions) compatible_stream: assumes \: "\ x y. x \ S \ x ~ y \ \ x \ \ y" assumes "pred_stream (\s. s \ S) xs" and [intro]: "x \ S" shows "pred_stream (\s. \ (reps (abss s)) = \ s) (x ## xs)" unfolding stream.pred_set proof clarify fix l u assume A: "(l, u) \ sset (x ## xs)" from assms have "pred_stream (\s. s \ S) (x ## xs)" by auto with A have "(l, u) \ S" by (fastforce simp: stream.pred_set) then have "abss (l, u) \ \" by auto then have "reps (abss (l, u)) ~ (l, u)" by simp with \ \(l, u) \ S\ show "\ (reps (abss (l, u))) = \ (l, u)" by blast qed lemma \_stream': "pred_stream (\s. \ (reps (abss s)) = \ s) (x ## xs)" if "pred_stream (\s. s \ S) xs" "x \ S" using compatible_stream[of \, OF \ that] . lemma \_stream': "pred_stream (\s. \ (reps (abss s)) = \ s) (x ## xs)" if "pred_stream (\s. s \ S) xs" "x \ S" using compatible_stream[of \, OF \ that] . lemmas \_stream = compatible_stream[of \, OF \] lemmas \_stream = compatible_stream[of \, OF \] subsection \Easier Result on All Configurations\ (* TODO: Rename *) lemma suntil_reps: assumes "\s\sset (smap abss y). s \ \" "(holds \' suntil holds \') (s' ## smap abss y)" shows "(holds \ suntil holds \) (s ## y)" using assms by (subst region_compatible_suntil[symmetric]; (intro \_stream \_stream)?) (auto simp: \'_def \'_def absp_def stream.pred_set \_abss_S s'_def comp_def) (* TODO: Rename *) lemma suntil_abss: assumes "\s\sset y. s \ S" "(holds \ suntil holds \) (s ## y)" shows "(holds \' suntil holds \') (s' ## smap abss y)" using assms by (subst (asm) region_compatible_suntil[symmetric]; (intro \_stream \_stream)?) (auto simp: \'_def \'_def absp_def stream.pred_set s'_def comp_def) (* TODO: Generalize to CTL formulae *) theorem P_sup_sunitl_eq: notes [measurable] = in_space_UNIV and [iff] = pred_stream_iff shows "(MDP.P_sup s (\x. (holds \ suntil holds \) (s ## x))) = (R_G.P_sup s' (\x. (holds \' suntil holds \') (s' ## x)))" unfolding MDP.P_sup_def R_G.P_sup_def proof (rule SUP_eq, goal_cases) case prems: (1 cfg) let ?cfg' = "absc cfg" from prems have "cfg \ valid_cfg" by (auto intro: MDP.valid_cfgI) then have "?cfg' \ R_G.valid_cfg" by (auto intro: R_G.valid_cfgI) from \cfg \ valid_cfg\ have alw_S: "almost_everywhere (MDP.T cfg) (pred_stream (\s. s \ S))" by (rule MDP.alw_S) from \?cfg'\ R_G.valid_cfg\ have alw_\: "almost_everywhere (R_G.T ?cfg') (pred_stream (\s. s \ \))" by (rule R_G.alw_S) have "emeasure (MDP.T cfg) {x \ space MDP.St. (holds \ suntil holds \) (s ## x)} = emeasure (R_G.T ?cfg') {x \ space R_G.St. (holds \' suntil holds \') (s' ## x)}" apply (rule path_measure_eq_absc1_new[symmetric, where P = "pred_stream (\ s. s \ \)" and Q = "pred_stream (\ s. s \ S)"] ) using prems alw_S alw_\ apply (auto intro: MDP.valid_cfgI simp: )[7] by (auto simp: S_abss_\ intro: \_abss_S intro!: suntil_abss suntil_reps, measurable) with prems show ?case by (inst_existentials ?cfg') auto next case prems: (2 cfg) let ?cfg' = "repcs s cfg" have "s = state ?cfg'" by simp from prems have "s' = state cfg" by auto have "pred_stream (\s. \ (reps (abss s)) = \ s) (state (repcs s cfg) ## x)" if "pred_stream (\s. s \ S) x" for x using prems that by (intro \_stream) auto moreover have "pred_stream (\s. \ (reps (abss s)) = \ s) (state (repcs s cfg) ## x)" if "pred_stream (\s. s \ S) x" for x using prems that by (intro \_stream) auto ultimately have "emeasure (R_G.T cfg) {x \ space R_G.St. (holds \' suntil holds \') (s' ## x)} = emeasure (MDP.T (repcs s cfg)) {x \ space MDP.St. (holds \ suntil holds \) (s ## x)}" apply (rewrite in "s ## _" \s = _\) apply (subst \s' = _\) unfolding \'_def \'_def s'_def apply (rule path_measure_eq_repcs''_new) using prems by (auto 4 3 simp: s'_def intro: R_G.valid_cfgI MDP.valid_cfgI) with prems show ?case by (inst_existentials ?cfg') auto qed end (* PTA Reachability Problem *) subsection \Divergent Adversaries\ context Probabilistic_Timed_Automaton begin definition "elapsed u u' \ Max ({u' c - u c | c. c \ \} \ {0})" definition "eq_elapsed u u' \ elapsed u u' > 0 \ (\ c \ \. u' c - u c = elapsed u u')" fun dur :: "('c, t) cval stream \ nat \ t" where "dur _ 0 = 0" | "dur (x ## y ## xs) (Suc i) = elapsed x y + dur (y ## xs) i" definition "divergent \ \ \ t. \ n. dur \ n > t" definition "div_cfg cfg \ AE \ in MDP.MC.T cfg. divergent (smap (snd o state) \)" definition "\_div \ \ \x \ \. (\ i. (\ j \ i. zero x (\ !! j)) \ (\ j \ i. \ zero x (\ !! j))) \ (\ i. \ j \ i. unbounded x (\ !! j))" definition "R_G_div_cfg cfg \ AE \ in MDP.MC.T cfg. \_div (smap (snd o state) \)" end context Probabilistic_Timed_Automaton_Regions begin definition "cfg_on_div st \ MDP.cfg_on st \ {cfg. div_cfg cfg}" definition "R_G_cfg_on_div st \ R_G.cfg_on st \ {cfg. R_G_div_cfg cfg}" lemma measurable_\_div[measurable]: "Measurable.pred MDP.MC.S \_div" unfolding \_div_def by (intro pred_intros_finite[OF beta_interp.finite] pred_intros_logic pred_intros_countable measurable_count_space_const measurable_compose[OF measurable_snth] ) measurable lemma elapsed_ge0[simp]: "elapsed x y \ 0" unfolding elapsed_def using finite(1) by auto lemma dur_pos: "dur xs i \ 0" apply (induction i arbitrary: xs) apply (auto; fail) subgoal for i xs apply (subst stream.collapse[symmetric]) apply (rewrite at "stl xs" stream.collapse[symmetric]) apply (subst dur.simps) by simp done lemma dur_mono: "i \ j \ dur xs i \ dur xs j" proof (induction i arbitrary: xs j) case 0 show ?case by (auto intro: dur_pos) next case (Suc i xs j) obtain x y ys where xs: "xs = x ## y ## ys" using stream.collapse by metis from Suc obtain j' where j': "j = Suc j'" by (cases j) auto with xs have "dur xs j = elapsed x y + dur (y ## ys) j'" by auto also from Suc j' have "\ \ elapsed x y + dur (y ## ys) i" by auto also have "elapsed x y + dur (y ## ys) i = dur xs (Suc i)" by (simp add: xs) finally show ?case . qed lemma dur_monoD: assumes "dur xs i < dur xs j" shows "i < j" using assms by - (rule ccontr; auto 4 4 dest: leI dur_mono[where xs = xs]) lemma elapsed_0D: assumes "c \ \" "elapsed u u' \ 0" shows "u' c - u c \ 0" proof - from assms have "u' c - u c \ {u' c - u c | c. c \ \} \ {0}" by auto with finite(1) have "u' c - u c \ Max ({u' c - u c | c. c \ \} \ {0})" by auto with assms(2) show ?thesis unfolding elapsed_def by auto qed lemma elapsed_ge: assumes "eq_elapsed u u'" "c \ \" shows "elapsed u u' \ u' c - u c" using assms unfolding eq_elapsed_def by (auto intro: elapsed_ge0 order.trans[OF elapsed_0D]) lemma elapsed_eq: assumes "eq_elapsed u u'" "c \ \" "u' c - u c \ 0" shows "elapsed u u' = u' c - u c" using elapsed_ge[OF assms(1,2)] assms unfolding eq_elapsed_def by auto lemma dur_shift: "dur \ (i + j) = dur \ i + dur (sdrop i \) j" apply (induction i arbitrary: \) apply simp subgoal for i \ apply simp apply (subst stream.collapse[symmetric]) apply (rewrite at "stl \" stream.collapse[symmetric]) apply (subst dur.simps) apply (rewrite in "dur \" stream.collapse[symmetric]) apply (rewrite in "dur (_ ## \) (Suc _)" stream.collapse[symmetric]) apply (subst dur.simps) apply simp done done lemma dur_zero: assumes "\ i. xs !! i \ \ !! i" "\ j \ i. zero x (\ !! j)" "x \ \" "\ i. eq_elapsed (xs !! i) (xs !! Suc i)" shows "dur xs i = 0" using assms proof (induction i arbitrary: xs \) case 0 then show ?case by simp next case (Suc i xs \) let ?x = "xs !! 0" let ?y = "xs !! 1" let ?ys = "stl (stl xs)" have xs: "xs = ?x ## ?y ## ?ys" by auto from Suc.prems have "\ i. (?y ## ?ys) !! i \ stl \ !! i" "\ j \ i. zero x (stl \ !! j)" "\ i. eq_elapsed (stl xs !! i) (stl xs !! Suc i)" by (metis snth.simps(2) | auto)+ from Suc.IH[OF this(1,2) \x \ _\] this(3) have [simp]: "dur (stl xs) i = 0" by auto from Suc.prems(1,2) have "?y x = 0" "?x x = 0" unfolding zero_def by force+ then have *: "?y x - ?x x = 0" by simp have "dur xs (Suc i) = elapsed ?x ?y" apply (subst xs) apply (subst dur.simps) by simp also have "\ = 0" apply (subst elapsed_eq[OF _ \x \ _\]) unfolding One_nat_def using Suc.prems(4) apply blast using * by auto finally show ?case . qed lemma dur_zero_tail: assumes "\ i. xs !! i \ \ !! i" "\ k \ i. k \ j \ zero x (\ !! k)" "x \ \" "j \ i" "\ i. eq_elapsed (xs !! i) (xs !! Suc i)" shows "dur xs j = dur xs i" proof - from \j \ i\ dur_shift[of xs i "j - i"] have "dur xs j = dur xs i + dur (sdrop i xs) (j - i)" by simp also have "\ = dur xs i" using assms by (rewrite in "dur (sdrop _ _) _" dur_zero[where \ = "sdrop i \"]) (auto dest: prop_nth_sdrop_pair[of eq_elapsed] prop_nth_sdrop prop_nth_sdrop_pair[of "(\)"]) finally show ?thesis . qed lemma elapsed_ge_pos: fixes u :: "('c, t) cval" assumes "eq_elapsed u u'" "c \ \" "u \ V" "u' \ V" shows "elapsed u u' \ u' c" proof (cases "elapsed u u' = 0") case True with assms show ?thesis by (auto simp: V_def) next case False from \u \ V\ \c \ \ \ have "u c \ 0" by (auto simp: V_def) from False assms have "elapsed u u' = u' c - u c" unfolding eq_elapsed_def by (auto simp add: less_le) also from \u c \ 0\ have "\ \ u' c" by simp finally show ?thesis . qed lemma dur_Suc: "dur xs (Suc i) - dur xs i = elapsed (xs !! i) (xs !! Suc i)" apply (induction i arbitrary: xs) apply simp apply (subst stream.collapse[symmetric]) apply (rewrite in "stl _" stream.collapse[symmetric]) apply (subst dur.simps) apply simp apply simp subgoal for i xs apply (subst stream.collapse[symmetric]) apply (rewrite in "stl _" stream.collapse[symmetric]) apply (subst dur.simps) apply simp apply (rewrite in "dur xs (Suc _)" stream.collapse[symmetric]) apply (rewrite at "stl xs" in "_ ## stl xs" stream.collapse[symmetric]) apply (subst dur.simps) apply simp done done inductive trans where succ: "t \ 0 \ u' = u \ t \ trans u u'" | reset: "set l \ \ \ u' = clock_set l 0 u \ trans u u'" | id: "u = u' \ trans u u'" abbreviation "stream_trans \ pairwise trans" lemma K_cfg_trans: assumes "cfg \ MDP.cfg_on (l, R)" "cfg' \ K_cfg cfg" "state cfg' = (l', R')" shows "trans R R'" using assms apply (simp add: set_K_cfg) apply (drule MDP.cfg_onD_action) apply (cases rule: K.cases) apply (auto intro: trans.intros) using admissible_targets_clocks(2) by (blast intro: trans.intros(2)) lemma enabled_stream_trans: assumes "cfg \ valid_cfg" "MDP.MC.enabled cfg xs" shows "stream_trans (smap (snd o state) xs)" using assms proof (coinduction arbitrary: cfg xs) case prems: (pairwise cfg xs) let ?xs = "stl (stl xs)" let ?x = "shd xs" let ?y = "shd (stl xs)" from MDP.pred_stream_cfg_on[OF prems] have *: "pred_stream (\cfg. state cfg \ S \ cfg \ MDP.cfg_on (state cfg)) xs" . obtain l R l' R' where eq: "state ?x = (l, R)" "state ?y = (l', R')" by force moreover from * have "?x \ MDP.cfg_on (state ?x)" "?x \ valid_cfg" by (auto intro: MDP.valid_cfgI simp: stream.pred_set) moreover from prems(2) have "?y \ K_cfg ?x" by (auto elim: MDP.MC.enabled.cases) ultimately have "trans R R'" by (intro K_cfg_trans[where cfg = ?x and cfg' = ?y and l = l and l' = l']) metis+ with \?x \ valid_cfg\ prems(2) show ?case apply (inst_existentials R R' "smap (snd o state) ?xs") apply (simp add: eq; fail)+ apply (rule disjI1, inst_existentials ?x "stl xs") by (auto simp: eq elim: MDP.MC.enabled.cases) qed lemma stream_trans_trans: assumes "stream_trans xs" shows "trans (xs !! i) (stl xs !! i)" using pairwise_Suc assms by auto lemma trans_eq_elapsed: assumes "trans u u'" "u \ V" shows "eq_elapsed u u'" using assms proof cases case (succ t) with finite(1) show ?thesis by (auto simp: cval_add_def elapsed_def max_def eq_elapsed_def) next case prems: (reset l) then have "u' c - u c \ 0" if "c \ \" for c using that \u \ V\ by (cases "c \ set l") (auto simp: V_def) then have "elapsed u u' = 0" unfolding elapsed_def using finite(1) apply simp apply (subst Max_insert2) by auto then show ?thesis by (auto simp: eq_elapsed_def) next case id then show ?thesis using finite(1) by (auto simp: Max_gr_iff elapsed_def eq_elapsed_def) qed lemma pairwise_trans_eq_elapsed: assumes "stream_trans xs" "pred_stream (\ u. u \ V) xs" shows "pairwise eq_elapsed xs" using trans_eq_elapsed assms by (auto intro: pairwise_mp simp: stream.pred_set) lemma not_reset_dur: assumes "\k>i. k \ j \ \ zero c ([xs !! k]\<^sub>\)" "j \ i" "c \ \" "stream_trans xs" "\ i. eq_elapsed (xs !! i) (xs !! Suc i)" "\ i. xs !! i \ V" shows "dur xs j - dur xs i = (xs !! j) c - (xs !! i) c" using assms proof (induction j) case 0 then show ?case by simp next case (Suc j) from stream_trans_trans[OF Suc.prems(4)] have trans: "trans (xs !! j) (xs !! Suc j)" by auto from Suc.prems have *: "\ zero c ([xs !! Suc j]\<^sub>\)" "eq_elapsed (xs !! j) (xs !! Suc j)" if "Suc j > i" using that by auto from Suc.prems(6) have "xs !! j \ V" "xs !! Suc j \ V" by blast+ then have regions: "[xs !! j]\<^sub>\ \ \" "[xs !! Suc j]\<^sub>\ \ \" by auto from trans have "(xs !! Suc j) c - (xs !! j) c \ 0" if "Suc j > i" proof (cases) case succ with regions show ?thesis by (auto simp: cval_add_def) next case prems: (reset l) show ?thesis proof (cases "c \ set l") case False with prems show ?thesis by auto next case True with prems have "(xs !! Suc j) c = 0" by auto moreover from assms have "xs !! Suc j \ [xs !! Suc j]\<^sub>\" by blast ultimately have "zero c ([xs !! Suc j]\<^sub>\)" using zero_all[OF finite(1) _ \c \ \\] regions(2) by (auto simp: \_def) with * that show ?thesis by auto qed next case id then show ?thesis by simp qed with * \c \ \\ elapsed_eq have *: "elapsed (xs !! j) (xs !! Suc j) = (xs !! Suc j) c - (xs !! j) c" if "Suc j > i" using that by blast show ?case proof (cases "i = Suc j") case False with Suc have "dur xs (Suc j) - dur xs i = dur xs (Suc j) - dur xs j + (xs !! j) c - (xs !! i) c" by auto also have "\ = elapsed (xs !! j) (xs !! Suc j) + (xs !! j) c - (xs !! i) c" by (simp add: dur_Suc) also have "\ = (xs !! Suc j) c - (xs !! j) c + (xs !! j) c - (xs !! i) c" using * False Suc.prems by auto also have "\ = (xs !! Suc j) c - (xs !! i) c" by simp finally show ?thesis by auto next case True then show ?thesis by simp qed qed lemma not_reset_dur': assumes "\j\i. \ zero c ([xs !! j]\<^sub>\)" "j \ i" "c \ \" "stream_trans xs" "\ i. eq_elapsed (xs !! i) (xs !! Suc i)" "\ j. xs !! j \ V" shows "dur xs j - dur xs i = (xs !! j) c - (xs !! i) c" using assms not_reset_dur by auto lemma not_reset_unbounded: assumes "\j\i. \ zero c ([xs !! j]\<^sub>\)" "j \ i" "c \ \" "stream_trans xs" "\ i. eq_elapsed (xs !! i) (xs !! Suc i)" "\ j. xs !! j \ V" "unbounded c ([xs !! i]\<^sub>\)" shows "unbounded c ([xs !! j]\<^sub>\)" proof - let ?u = "xs !! i" let ?u' = "xs !! j" let ?R = "[xs !! i]\<^sub>\" from assms have "?u \ ?R" by auto from assms(6) have "?R \ \" by auto then obtain I r where "?R = region \ I r" "valid_region \ k I r" unfolding \_def by auto with assms(3,7) unbounded_Greater \?u \ ?R\ have "?u c > k c" by force also from not_reset_dur'[OF assms(1-6)] dur_mono[OF \j \ i\, of xs] have "?u' c \ ?u c" by auto finally have "?u' c > k c" by auto let ?R' = "[xs !! j]\<^sub>\" from assms have "?u' \ ?R'" by auto from assms(6) have "?R' \ \" by auto then obtain I r where "?R' = region \ I r" "valid_region \ k I r" unfolding \_def by auto moreover with \?u' c > _\ \?u' \ _\ gt_GreaterD \c \ \\ have "I c = Greater (k c)" by auto ultimately show ?thesis using Greater_unbounded[OF finite(1) _ \c \ \\] by auto qed lemma gt_unboundedD: assumes "u \ R" and "R \ \" and "c \ \" and "real (k c) < u c" shows "unbounded c R" proof - from assms obtain I r where "R = region \ I r" "valid_region \ k I r" unfolding \_def by auto with Greater_unbounded[of \ k I r c] gt_GreaterD[of u \ I r k c] assms finite(1) show ?thesis by auto qed definition trans' :: "('c, t) cval \ ('c, t) cval \ bool" where "trans' u u' \ ((\ c \ \. u c > k c \ u' c > k c \ u \ u') \ u' = u \ 0.5) \ ((\ c \ \. u c = 0 \ u' c > 0 \ (\c\\. \d. d \ k c \ u' c = real d)) \ u' = delayedR ([u']\<^sub>\) u)" (* XXX Move *) lemma zeroI: assumes "c \ \" "u \ V" "u c = 0" shows "zero c ([u]\<^sub>\)" proof - from assms have "u \ [u]\<^sub>\" "[u]\<^sub>\ \ \" by auto then obtain I r where "[u]\<^sub>\ = region \ I r" "valid_region \ k I r" unfolding \_def by auto with zero_all[OF finite(1) this(2) \c \ \\] \u \ [u]\<^sub>\\ \u c = 0\ show ?thesis by auto qed (* XXX Move, rename *) lemma zeroD: "u x = 0" if "zero x ([u]\<^sub>\)" "u \ V" using that by (metis regions_part_ex(1) zero_def) lemma not_zeroD: assumes "\ zero x ([u]\<^sub>\)" "u \ V" "x \ \" shows "u x > 0" proof - from zeroI assms have "u x \ 0" by auto moreover from assms have "u x \ 0" unfolding V_def by auto ultimately show ?thesis by auto qed (* XXX Move *) lemma not_const_intv: assumes "u \ V" "\c\\. \d. d \ k c \ u c = real d" shows "\c\\. \u \ [u]\<^sub>\. \d. d \ k c \ u c = real d" proof - from assms have "u \ [u]\<^sub>\" "[u]\<^sub>\ \ \" by auto then obtain I r where I: "[u]\<^sub>\ = region \ I r" "valid_region \ k I r" unfolding \_def by auto have "\d. d \ k c \ u' c = real d" if "c \ \" "u' \ [u]\<^sub>\" for c u' proof safe fix d assume A: "d \ k c" "u' c = real d" from I that have "intv_elem c u' (I c)" "valid_intv (k c) (I c)" by auto then show False using A I \u \ [u]\<^sub>\\ \c \ \\ assms(2) by (cases; fastforce) qed then show ?thesis by auto qed lemma K_cfg_trans': assumes "repcs (l, u) cfg \ MDP.cfg_on (l, u)" "cfg' \ K_cfg (repcs (l, u) cfg)" "state cfg' = (l', u')" "(l, u) \ S" "cfg \ R_G.valid_cfg" "abss (l, u) = state cfg" shows "trans' u u'" using assms apply (simp add: set_K_cfg) apply (drule MDP.cfg_onD_action) apply (cases rule: K.cases) apply assumption proof goal_cases case prems: (1 l u t) from assms \_ = (l, u)\ have "repcs (l, u) cfg \ valid_cfg" by (auto intro: MDP.valid_cfgI) then have "absc (repcs (l, u) cfg) \ R_G.valid_cfg" by auto from prems have *: "rept (l, u) (action cfg) = return_pmf (l, u \ t)" unfolding repcs_def by auto from \abss _ = _\ \_ = (l, u)\ \cfg \ R_G.valid_cfg\ have "action cfg \ \ (abss (l, u))" by (auto dest: R_G_I) from abst_rept_id[OF this] * have "action cfg = abst (return_pmf (l, u \ t))" by auto with prems have **: "action cfg = return_pmf (l, [u \ t]\<^sub>\)" unfolding abst_def by auto show ?thesis proof (cases "\ c \ \. u c > k c") case True from prems have "u \ t \ [u]\<^sub>\" by (auto intro: upper_right_closed[OF True]) with prems have "[u \ t]\<^sub>\ = [u]\<^sub>\" by (auto dest: alpha_interp.region_unique_spec) with ** have "action cfg = return_pmf (l, [u]\<^sub>\)" by simp with True have "rept (l, u) (action cfg) = return_pmf (l, u \ 0.5)" unfolding rept_def using prems by auto with * have "u \ t = u \ 0.5" by auto moreover from prems have "u' = u \ t" by auto moreover from prems True have "\ c \ \. u' c > k c" by (auto simp: cval_add_def) ultimately show ?thesis using True \_ = (l, u)\ unfolding trans'_def by auto next case F: False show ?thesis proof (cases "\c\\. u c = 0 \ 0 < u' c \ (\c\\. \d. d \ k c \ u' c = real d)") case True from prems have "u' \ [u']\<^sub>\" by auto from prems have "[u \ t]\<^sub>\ \ Succ \ ([u]\<^sub>\)" by auto from True obtain c where "c \ \" "u c = 0" "u' c > 0" by auto with zeroI prems have "zero c ([u]\<^sub>\)" by auto moreover from \u' \ _\ \u' c > 0\ have "\ zero c ([u']\<^sub>\)" unfolding zero_def by fastforce ultimately have "[u \ t]\<^sub>\ \ [u]\<^sub>\" using prems by auto moreover from True not_const_intv prems have "\ u \ [u \ t]\<^sub>\. \c\\. \d. d \ k c \ u c = real d" by auto ultimately have "\R'. (l, u) \ S \ action cfg = return_pmf (l, R') \ R' \ Succ \ ([u]\<^sub>\) \ [u]\<^sub>\ \ R' \ (\u\R'. \c\\. \d. d \ k c \ u c = real d)" apply - apply (rule exI[where x = "[u \ t]\<^sub>\"]) apply safe using prems ** by auto then have "rept (l, u) (action cfg) = return_pmf (l, delayedR (SOME R'. action cfg = return_pmf (l, R')) u)" unfolding rept_def by auto with * ** prems have "u' = delayedR ([u \ t]\<^sub>\) u" by auto with F True prems show ?thesis unfolding trans'_def by auto next case False with F \_ = (l, u)\ show ?thesis unfolding trans'_def by auto qed qed next case prems: (2 _ _ \ \) then obtain X where X: "u' = ([X := 0]u)" "(X, l') \ set_pmf \" by auto from \_ \ S\ have "u \ V" by auto let ?r = "SOME r. set r = X" show ?case proof (cases "X = {}") case True with X have "u = u'" by auto with non_empty show ?thesis unfolding trans'_def by auto next case False then obtain x where "x \ X" by auto moreover have "X \ \" using admissible_targets_clocks(1)[OF prems(10) X(2)] by auto ultimately have "x \ \" by auto from \X \ \\ finite(1) obtain r where "set r = X" using finite_list finite_subset by blast then have r: "set ?r = X" by (rule someI) with \x \ X\ X have "u' x = 0" by auto from X r \u \ V\ \X \ \\ have "u' x \ u x" for x by (cases "x \ X"; auto simp: V_def) have False if "u' x > 0 \ u x = 0" for x using \u' _ \ _\[of x] that by auto with \u' x = 0\ show ?thesis using \x \ \\ unfolding trans'_def by auto qed next case 3 with non_empty show ?case unfolding trans'_def by auto qed coinductive enabled_repcs where "enabled_repcs (shd xs) (stl xs) \ shd xs = repcs st' cfg' \ st' \ rept st (action cfg) \ abss st' = state cfg' \ cfg' \ R_G.valid_cfg \ enabled_repcs (repcs st cfg) xs" (* XXX Move *) lemma K_cfg_rept_in: assumes "cfg \ R_G.valid_cfg" and "abss st = state cfg" and "cfg' \ K_cfg cfg" shows "(THE s'. s' \ set_pmf (rept st (action cfg)) \ abss s' = state cfg') \ set_pmf (rept st (action cfg))" proof - from assms(1,2) have "action cfg \ \ (abss st)" by (auto simp: R_G_I) from \cfg' \ _\ have "cfg' = cont cfg (state cfg')" "state cfg' \ action cfg" by (auto simp: set_K_cfg) with abst_rept_id[OF \action _ \ _\] pmf.set_map have "state cfg' \ abss ` set_pmf (rept st (action cfg))" unfolding abst_def by metis then obtain st' where "st' \ rept st (action cfg)" "abss st' = state cfg'" unfolding abst_def by auto with K_cfg_rept_aux[OF assms(1,2) this(1)] show ?thesis by auto qed lemma enabled_repcsI: assumes "cfg \ R_G.valid_cfg" "abss st = state cfg" "MDP.MC.enabled (repcs st cfg) xs" shows "enabled_repcs (repcs st cfg) xs" using assms proof (coinduction arbitrary: cfg xs st) case prems: (enabled_repcs cfg xs st) let ?x = "shd xs" and ?y = "shd (stl xs)" let ?st = "THE s'. s' \ set_pmf (rept st (action cfg)) \ abss s' = state (absc ?x)" from prems(3) have "?x \ K_cfg (repcs st cfg)" by cases with K_cfg_map_repcs[OF prems(1,2)] obtain cfg' where "cfg' \ K_cfg cfg" "?x = repcs (THE s'. s' \ rept st (action cfg) \ abss s' = state cfg') cfg'" by auto let ?st = "THE s'. s' \ rept st (action cfg) \ abss s' = state cfg'" from K_cfg_rept_action[OF prems(1,2) \cfg' \ _\] have "abss ?st = state cfg'" . moreover from K_cfg_rept_in[OF prems(1,2) \cfg' \ _\] have "?st \ rept st (action cfg)" . moreover have "cfg' \ R_G.valid_cfg" using \cfg' \ K_cfg cfg\ prems(1) by blast moreover from absc_repcs_id[OF this \abss ?st = state cfg'\] \?x = _\ have "absc ?x = cfg'" by auto moreover from prems(3) have "MDP.MC.enabled (shd xs) (stl xs)" by cases ultimately show ?case using \?x = _\ by (inst_existentials xs ?st "absc ?x" st cfg) fastforce+ qed lemma repcs_eq_rept: "rept st (action cfg) = rept st'' (action cfg'')" if "repcs st cfg = repcs st'' cfg''" by (metis (mono_tags, lifting) action_cfg_corec old.prod.case repcs_def that) lemma enabled_stream_trans': assumes "cfg \ R_G.valid_cfg" "abss st = state cfg" "MDP.MC.enabled (repcs st cfg) xs" shows "pairwise trans' (smap (snd o state) xs)" using assms proof (coinduction arbitrary: cfg xs st) case prems: (pairwise cfg xs) let ?xs = "stl xs" from prems have A: "enabled_repcs (repcs st cfg) xs" by (auto intro: enabled_repcsI) then obtain st' cfg' where "enabled_repcs (shd xs) (stl xs)" "shd xs = repcs st' cfg'" "st' \ rept st (action cfg)" "abss st' = state cfg'" "cfg' \ R_G.valid_cfg" apply atomize_elim apply (cases rule: enabled_repcs.cases) apply assumption subgoal for st' cfg' st'' cfg'' by (inst_existentials st' cfg') (auto dest: repcs_eq_rept) done then obtain st'' cfg'' where "enabled_repcs (shd ?xs) (stl ?xs)" "shd ?xs = repcs st'' cfg''" "st'' \ rept st' (action cfg')" "abss st'' = state cfg''" by atomize_elim (subst (asm)enabled_repcs.simps, fastforce dest: repcs_eq_rept) let ?x = "shd xs" let ?y = "shd (stl xs)" let ?cfg = "repcs st cfg" from prems have "?cfg \ valid_cfg" by auto from MDP.pred_stream_cfg_on[OF \?cfg \ valid_cfg\ prems(3)] have *: "pred_stream (\cfg. state cfg \ S \ cfg \ MDP.cfg_on (state cfg)) xs" . obtain l u l' u' where eq: "st' = (l, u)" "st'' = (l', u')" by force moreover from * have "?x \ MDP.cfg_on (state ?x)" "?x \ valid_cfg" by (auto intro: MDP.valid_cfgI simp: stream.pred_set) moreover from prems(3) have "?y \ K_cfg ?x" by (auto elim: MDP.MC.enabled.cases) ultimately have "trans' u u'" using \?x = _\ \?y = _\ \cfg' \ _\ \abss st' = _\ by (intro K_cfg_trans') (auto dest: MDP.valid_cfg_state_in_S) with \?x \ valid_cfg\ \cfg' \ R_G.valid_cfg\ prems(3) \abss _ = state cfg'\ show ?case apply (inst_existentials u u' "smap (snd o state) (stl ?xs)") apply (simp add: eq \?x = _\ \?y = _\; fail)+ by ((intro disjI1 exI)?; auto simp: \?x = _\ \?y = _\ eq elim: MDP.MC.enabled.cases) qed lemma divergent_\_divergent: assumes in_S: "pred_stream (\ u. u \ V) xs" and div: "divergent xs" and trans: "stream_trans xs" shows "\_div (smap (\ u. [u]\<^sub>\) xs)" (is "\_div ?\") unfolding \_div_def proof (safe, simp_all) fix x i assume x: "x \ \" and bounded: "\i. \j\i. \ unbounded x ([xs !! j]\<^sub>\)" from in_S have xs_\: "\i. xs !! i \ ?\ !! i" by (auto simp: stream.pred_set) from trans in_S have elapsed: "\ i. eq_elapsed (xs !! i) (xs !! Suc i)" by (fastforce intro: pairwise_trans_eq_elapsed pairwise_Suc[where P = eq_elapsed]) { assume A: "\j \ i. \ zero x ([xs !! j]\<^sub>\)" let ?t = "dur xs i + k x" from div obtain j where j: "dur xs j > dur xs i + k x" unfolding divergent_def by auto then have "k x < dur xs j - dur xs i" by auto also with not_reset_dur'[OF A less_imp_le[OF dur_monoD], of xs] \x \ \\ assms elapsed have "\ = (xs !! j) x - (xs !! i) x" by (auto simp: stream.pred_set) also have "\ \ (xs !! j) x" using assms(1) \x \ \\ unfolding V_def by (auto simp: stream.pred_set) finally have "unbounded x ([xs !! j]\<^sub>\)" using assms \x \ \\ by (intro gt_unboundedD) (auto simp: stream.pred_set) moreover from dur_monoD[of xs i j] j A have "\j' \ j. \ zero x ([xs !! j']\<^sub>\)" by auto ultimately have "\i\j. unbounded x ([xs !! i]\<^sub>\)" using elapsed assms x by (auto intro: not_reset_unbounded simp: stream.pred_set) with bounded have False by auto } then show "\j\i. zero x ([xs !! j]\<^sub>\)" by auto { assume A: "\j \ i. zero x ([xs !! j]\<^sub>\)" from div obtain j where j: "dur xs j > dur xs i" unfolding divergent_def by auto then have "j \ i" by (auto dest: dur_monoD) from A have "\j \ i. zero x (?\ !! j)" by auto with dur_zero_tail[OF xs_\ _ x \i \ j\ elapsed] j have False by simp } then show "\j\i. \ zero x ([xs !! j]\<^sub>\)" by auto qed lemma (in -) fixes f :: "nat \ real" assumes "\ i. f i \ 0" "\ i. \ j \ i. f j > d" "d > 0" shows "\ n. (\ i \ n. f i) > t" oops (* TODO: Reduce this proof to a more general theorem *) lemma dur_ev_exceedsI: assumes "\ i. \ j \ i. dur xs j - dur xs i \ d" and "d > 0" obtains i where "dur xs i > t" proof - have base: "\ i. dur xs i > t" if "t < d" for t proof - from assms obtain j where "dur xs j - dur xs 0 \ d" by fastforce with dur_pos[of xs 0] have "dur xs j \ d" by simp with \d > 0\ \t < d\ show ?thesis by - (rule exI[where x = j]; auto) qed have base2: "\ i. dur xs i > t" if "t \ d" for t proof (cases "t = d") case False with \t \ d\ base show ?thesis by simp next case True from base \d > 0\ obtain i where "dur xs i > 0" by auto moreover from assms obtain j where "dur xs j - dur xs i \ d" by auto ultimately have "dur xs j > d" by auto with \t = d\ show ?thesis by auto qed show ?thesis proof (cases "t \ 0") case False with dur_pos have "dur xs 0 > t" by auto then show ?thesis by (fastforce intro: that) next case True let ?m = "nat \t / d\" from True have "\ i. dur xs i > ?m * d" proof (induction ?m arbitrary: t) case 0 with base[OF \0 < d\] show ?case by simp next case (Suc n t) let ?t = "t - d" show ?case proof (cases "t \ d") case True have "?t / d = t / d - 1" (* Generated by sledgehammer *) (* Alternative: by (smt assms(2) diff_divide_distrib divide_self_if) *) proof - have "t / d + - 1 * ((t + - 1 * d) / d) + - 1 * (d / d) = 0" by (simp add: diff_divide_distrib) then have "t / d + - 1 * ((t + - 1 * d) / d) = 1" using assms(2) by fastforce then show ?thesis by algebra qed then have "\?t / d\ = \t / d\ - 1" by simp with \Suc n = _\ have "n = nat \?t / d\" by simp with Suc \t \ d\ obtain i where "nat \?t / d\ * d < dur xs i" by fastforce from assms obtain j where "dur xs j - dur xs i \ d" "j \ i" by auto with \dur xs i > _\ have "nat \?t / d\ * d + d < dur xs j" by simp with True have "dur xs j > nat \t / d\ * d" by (metis Suc.hyps(2) \n = nat \(t - d) / d\\ add.commute distrib_left mult.commute mult.right_neutral of_nat_Suc) then show ?thesis by blast next case False with \t \ 0\ \d > 0\ have "nat \t / d\ \ 1" by simp then have "nat \t / d\ * d \ d" by (metis One_nat_def \Suc n = _\ Suc_leI add.right_neutral le_antisym mult.commute mult.right_neutral of_nat_0 of_nat_Suc order_refl zero_less_Suc) with base2 show ?thesis by auto qed qed then obtain i where "dur xs i > ?m * d" by atomize_elim moreover from \t \ 0\ \d > 0\ have "?m * d \ t" using pos_divide_le_eq real_nat_ceiling_ge by blast ultimately show ?thesis using that[of i] by simp qed qed lemma not_reset_mono: assumes "stream_trans xs" "shd xs c1 \ shd xs c2" "stream_all (\ u. u \ V) xs" "c2 \ \" shows "(holds (\ u. u c1 \ u c2) until holds (\ u. u c1 = 0)) xs" using assms proof (coinduction arbitrary: xs) case prems: (UNTIL xs) let ?xs = "stl xs" let ?x = "shd xs" let ?y = "shd ?xs" show ?case proof (cases "?x c1 = 0") case False show ?thesis proof (cases "?y c1 = 0") case False from prems have "trans ?x ?y" by (intro pairwise_pairD[of trans]) then have "?y c1 \ ?y c2" proof cases case A: (reset t) show ?thesis proof (cases "c1 \ set t") case True with A False show ?thesis by auto next case False from prems have "?x c2 \ 0" by (auto simp: V_def) with A have "?y c2 \ ?x c2" by (cases "c2 \ set t") auto with A False \?x c1 \ ?x c2\ show ?thesis by auto qed qed (use prems in \auto simp: cval_add_def\) moreover from prems have "stream_trans ?xs" "stream_all (\ u. u \ V) ?xs" by (auto intro: pairwise_stlD stl_sset) ultimately show ?thesis using prems by auto qed (use prems in \auto intro: UNTIL.base\) qed auto qed lemma \_divergent_divergent_aux: fixes xs :: "('c, t) cval stream" assumes "stream_trans xs" "stream_all (\ u. u \ V) xs" "(xs !! i) c1 = 0" "\ k > i. k \ j \ (xs !! k) c2 = 0" "\ k > i. k \ j \ (xs !! k) c1 \ 0" "c1 \ \" "c2 \ \" shows "(xs !! j) c1 \ (xs !! j) c2" proof - from assms obtain k where k: "k > i" "k \ j" "(xs !! k) c2 = 0" by auto with assms(5) \k \ j\ have "(xs !! k) c1 \ 0" by auto moreover from assms(2) \c1 \ \\ have "(xs !! k) c1 \ 0" by (auto simp: V_def) ultimately have "(xs !! k) c1 > 0" by auto with \(xs !! k) c2 = 0\ have "shd (sdrop k xs) c1 \ shd (sdrop k xs) c2" by auto from not_reset_mono[OF _ this] assms have "(holds (\u. u c2 \ u c1) until holds (\u. u c1 = 0)) (sdrop k xs)" by (auto intro: sset_sdrop pairwise_sdropD) from assms(5) k(2) \k > i\ have "\ m \ j - k. (sdrop k xs !! m) c1 \ 0" by simp with holds_untilD[OF \(_ until _) _\, of "j - k"] have "(sdrop k xs !! (j - k)) c2 \ (sdrop k xs !! (j - k)) c1" . then show "(xs !! j) c2 \ (xs !! j) c1" using k(1,2) by simp qed lemma unbounded_all: assumes "R \ \" "u \ R" "unbounded x R" "x \ \" shows "u x > k x" proof - from assms obtain I r where R: "R = region \ I r" "valid_region \ k I r" unfolding \_def by auto with unbounded_Greater \x \ \\ assms(3) have "I x = Greater (k x)" by simp with \u \ R\ R \x \ \\ show ?thesis by force qed lemma trans_not_delay_mono: "u' c \ u c" if "trans u u'" "u \ V" "x \ \" "u' x = 0" "c \ \" using \trans u u'\ proof (cases) case (reset l) with that show ?thesis by (cases "c \ set l") (auto simp: V_def) qed (use that in \auto simp: cval_add_def V_def add_nonneg_eq_0_iff\) lemma dur_reset: assumes "pairwise eq_elapsed xs" "pred_stream (\ u. u \ V) xs" "zero x ([xs !! Suc i]\<^sub>\)" "x \ \" shows "dur xs (Suc i) - dur xs i = 0" proof - from assms(2) have in_V: "xs !! Suc i \ V" unfolding stream.pred_set by auto (metis snth.simps(2) snth_sset) with elapsed_ge_pos[of "xs !! i" "xs !! Suc i" x] pairwise_Suc[OF assms(1)] assms(2-) have "elapsed (xs !! i) (xs !! Suc i) \ (xs !! Suc i) x" unfolding stream.pred_set by auto with in_V assms(3) have "elapsed (xs !! i) (xs !! Suc i) \ 0" by (auto simp: zeroD) with elapsed_ge0[of "xs !! i" "xs !! Suc i"] have "elapsed (xs !! i) (xs !! Suc i) = 0" by linarith then show ?thesis by (subst dur_Suc) qed lemma resets_mono_0': assumes "pairwise eq_elapsed xs" "stream_all (\ u. u \ V) xs" "stream_trans xs" "\ j \ i. zero x ([xs !! j]\<^sub>\)" "x \ \" "c \ \" shows "(xs !! i) c = (xs !! 0) c \ (xs !! i) c = 0" using assms proof (induction i) case 0 then show ?case by auto next case (Suc i) from Suc.prems have *: "(xs !! Suc i) x = 0" "(xs !! i) x = 0" by (blast intro: zeroD snth_sset, force intro: zeroD snth_sset) from pairwise_Suc[OF Suc.prems(3)] have "trans (xs !! i) (xs !! Suc i)" . then show ?case proof cases case prems: (succ t) with * have "t = 0" unfolding cval_add_def by auto with prems have "(xs !! Suc i) c = (xs !! i) c" unfolding cval_add_def by auto with Suc show ?thesis by auto next case prems: (reset l) then have "(xs !! Suc i) c = 0 \ (xs !! Suc i) c = (xs !! i) c" by (cases "c \ set l") auto with Suc show ?thesis by auto next case id with Suc show ?thesis by auto qed qed lemma resets_mono': assumes "pairwise eq_elapsed xs" "pred_stream (\ u. u \ V) xs" "stream_trans xs" "\ k \ i. k \ j \ zero x ([xs !! k]\<^sub>\)" "x \ \" "c \ \" "i \ j" shows "(xs !! j) c = (xs !! i) c \ (xs !! j) c = 0" using assms proof - from assms have 1: "stream_all (\ u. u \ V) (sdrop i xs)" using sset_sdrop unfolding stream.pred_set by force from assms have 2: "pairwise eq_elapsed (sdrop i xs)" by (intro pairwise_sdropD) from assms have 3: "stream_trans (sdrop i xs)" by (intro pairwise_sdropD) from assms have 4: "\k\j - i. zero x ([sdrop i xs !! k]\<^sub>\)" - by (simp add: Nat.le_diff_conv2 assms(6)) + by (simp add: le_diff_conv2 assms(6)) from resets_mono_0'[OF 2 1 3 4 assms(5,6)] \i \ j\ show ?thesis by simp qed lemma resets_mono: assumes "pairwise eq_elapsed xs" "pred_stream (\ u. u \ V) xs" "stream_trans xs" "\ k \ i. k \ j \ zero x ([xs !! k]\<^sub>\)" "x \ \" "c \ \" "i \ j" shows "(xs !! j) c \ (xs !! i) c" using assms using assms by (auto simp: V_def dest: resets_mono'[where c = c] simp: stream.pred_set) lemma \_divergent_divergent_aux2: fixes M :: "(nat \ bool) set" assumes "\ i. \ P \ M. \ j \ i. P j" "M \ {}" "finite M" shows "\i.\j\i.\k>j.\ P \ M. P j \ P k \ (\ m < k. j < m \ \ P m) \ (\ Q \ M. \ m \ k. j < m \ Q m)" proof fix i let ?j1 = "Max {LEAST m. m > i \ P m | P. P \ M}" from \M \ {}\ obtain P where "P \ M" by auto let ?m = "LEAST m. m > i \ P m" from assms(1) \P \ M\ obtain j where "j \ Suc i" "P j" by auto then have "j > i" "P j" by auto with \P \ M\ have "?m > i \ P ?m" by - (rule LeastI; auto) moreover with \finite M\ \P \ M\ have "?j1 \ ?m" by - (rule Max_ge; auto) ultimately have "?j1 \ i" by simp moreover have "\ m > i. m \ ?j1 \ P m" if "P \ M" for P proof - let ?m = "LEAST m. m > i \ P m" from assms(1) \P \ M\ obtain j where "j \ Suc i" "P j" by auto then have "j > i" "P j" by auto with \P \ M\ have "?m > i \ P ?m" by - (rule LeastI; auto) moreover with \finite M\ \P \ M\ have "?j1 \ ?m" by - (rule Max_ge; auto) ultimately show ?thesis by auto qed ultimately obtain j1 where j1: "j1 \ i" "\ P \ M. \ m > i. j1 \ m \ P m" by auto define k where "k Q = (LEAST k. k > j1 \ Q k)" for Q let ?k = "Max {k Q | Q. Q \ M}" let ?P = "SOME P. P \ M \ k P = ?k" let ?j = "Max {j. i \ j \ j \ j1 \ ?P j}" have "?k \ {k Q | Q. Q \ M}" using assms by - (rule Max_in; auto) then obtain P where P: "k P = ?k" "P \ M" by auto have "?k \ k Q" if "Q \ M" for Q using assms that by - (rule Max_ge; auto) have *: "?P \ M \ k ?P = ?k" using P by - (rule someI[where x = P]; auto) with j1 have "\ m > i. j1 \ m \ ?P m" by auto with \finite _\ have "?j \ {j. i \ j \ j \ j1 \ ?P j}" by - (rule Max_in; auto) have k: "k Q > j1 \ Q (k Q)" if "Q \ M" for Q proof - from assms(1) \Q \ M\ obtain m where "m \ Suc j1" "Q m" by auto then have "m > j1" "Q m" by auto then show "k Q > j1 \ Q (k Q)" unfolding k_def by - (rule LeastI; blast) qed with * \?j \ _\ have "?P ?k" "?j < ?k" by fastforce+ have "\ ?P m" if "?j < m" "m < ?k" for m proof (rule ccontr, simp) assume "?P m" have "m > j1" proof (rule ccontr) assume "\ j1 < m" with \?j < m\ \?j \ _\ have "i \ m" "m \ j1" by auto with \?P m\ \finite _\ have "?j \ m" by - (rule Max_ge; auto) with \?j < m\ show False by simp qed with \?P m\ \finite _\ have "k ?P \ m" unfolding k_def by (auto intro: Least_le) with * \m < ?k\ show False by auto qed moreover have "\ m \ ?k. ?j < m \ Q m" if "Q \ M" for Q proof - from k[OF \Q \ M\] have "k Q > j1 \ Q (k Q)" . moreover with \finite _\ \Q \ M\ have "k Q \ ?k" by - (rule Max_ge; auto) moreover with \?j \ _\ \k Q > _ \ _\ have "?j < k Q" by auto ultimately show ?thesis by auto qed ultimately show "\j\i.\k>j.\ P \ M. P j \ P k \ (\ m < k. j < m \ \ P m) \ (\ Q \ M. \ m \ k. j < m \ Q m)" using \?j < ?k\ \?j \ _\ \?P ?k\ * by (inst_existentials ?j ?k ?P; blast) qed lemma \_divergent_divergent: assumes in_S: "pred_stream (\ u. u \ V) xs" and div: "\_div (smap (\ u. [u]\<^sub>\) xs)" and trans: "stream_trans xs" and trans': "pairwise trans' xs" and unbounded_not_const: "\u. (\c\\. real (k c) < u c) \ \ ev (alw (\xs. shd xs = u)) xs" shows "divergent xs" unfolding divergent_def proof fix t from pairwise_trans_eq_elapsed[OF trans in_S] have eq_elapsed: "pairwise eq_elapsed xs" . define X1 where "X1 = {x. x \ \ \ (\i. \ j \ i. unbounded x ([xs !! j]\<^sub>\))}" let ?i = "Max {(SOME i. \ j \ i. unbounded x ([xs !! j]\<^sub>\)) | x. x \ \}" from finite(1) non_empty have "?i \ {(SOME i. \ j \ i. unbounded x ([xs !! j]\<^sub>\)) | x. x \ \}" by (intro Max_in) auto have "unbounded x ([xs !! j]\<^sub>\)" if "x \ X1" "j \ ?i" for x j proof - have "X1 \ \" unfolding X1_def by auto with finite(1) non_empty \x \ X1\ have *: "?i \ (SOME i. \ j \ i. unbounded x ([xs !! j]\<^sub>\))" (is "?i \ ?k") by (intro Max_ge) auto from \x \ X1\ have "\ k. \ j \ k. unbounded x ([xs !! j]\<^sub>\)" by (auto simp: X1_def) then have "\ j \ ?k. unbounded x ([xs !! j]\<^sub>\)" by (rule someI_ex) moreover from \j \ ?i\ \?i \ _\ have "j \ ?k" by auto ultimately show ?thesis by blast qed then obtain i where unbounded: "\ x \ X1. \ j \ i. unbounded x ([xs !! j]\<^sub>\)" using finite by auto show "\ n. t < dur xs n" proof (cases "\ x \ \. (\i. \ j \ i. unbounded x ([xs !! j]\<^sub>\))") case True then have "X1 = \" unfolding X1_def by auto have "\k\j. 0.5 \ dur xs k - dur xs j" for j proof - let ?u = "xs !! max i j" from in_S have "?u \ [?u]\<^sub>\" "[?u]\<^sub>\ \ \" by (auto simp: stream.pred_set) moreover from unbounded \X1 = \\ have "\ x \ \. unbounded x ([?u]\<^sub>\)" by force ultimately have "\ x \ \. ?u x > k x" by (auto intro: unbounded_all) with unbounded_not_const have "\ ev (alw (HLD {?u})) xs" unfolding HLD_iff by simp then obtain r where "r \ max i j" "xs !! r \ xs !! Suc r" apply atomize_elim apply (simp add: not_ev_iff not_alw_iff) apply (drule alw_sdrop[where n = "max i j"]) apply (drule alwD) apply (subst (asm) (3) stream.collapse[symmetric]) apply simp apply (drule ev_neq_start_implies_ev_neq[simplified comp_def]) using stream.collapse[of "sdrop (max i j) xs"] by (auto 4 3 elim: ev_sdropD) let ?k = "Suc r" from in_S have "xs !! ?k \ V" using snth_sset unfolding stream.pred_set by blast with in_S have *: "xs !! r \ [xs !! r]\<^sub>\" "[xs !! r]\<^sub>\ \ \" "xs !! ?k \ [xs !! ?k]\<^sub>\" "[xs !! ?k]\<^sub>\ \ \" by (auto simp: stream.pred_set) from \r \ _\ have "r \ i" "?k \ i" by auto with unbounded \X1 = \\ have "\ x \ \. unbounded x ([xs !! r]\<^sub>\)" "\ x \ \. unbounded x ([xs !! ?k]\<^sub>\)" by (auto simp del: snth.simps(2)) with in_S have "\ x \ \. (xs !! r) x > k x" "\ x \ \. (xs !! ?k) x > k x" using * by (auto intro: unbounded_all) moreover from trans' have "trans' (xs !! r) (xs !! ?k)" using pairwise_Suc by auto ultimately have "(xs !! ?k) = (xs !! r) \ 0.5" unfolding trans'_def using \xs !! r \ _\ by auto moreover from pairwise_Suc[OF eq_elapsed] have "eq_elapsed (xs !! r) (xs !! ?k)" by auto ultimately have "dur xs ?k - dur xs r = 0.5" using non_empty by (auto simp: cval_add_def dur_Suc elapsed_eq) with dur_mono[of j r xs] \r \ max i j\ have "dur xs ?k - dur xs j \ 0.5" by auto with \r \ max i j\ show ?thesis by - (rule exI[where x = ?k]; auto) qed then show ?thesis by - (rule dur_ev_exceedsI[where d = "0.5"]; auto) next case False define X2 where "X2 = \ - X1" from False have "X2 \ {}" unfolding X1_def X2_def by fastforce have inf_resets: "\i. (\j\i. zero x ([xs !! j]\<^sub>\)) \ (\j\i. \ zero x ([xs !! j]\<^sub>\))" if "x \ X2" for x using that div unfolding X1_def X2_def \_div_def by fastforce have "\ j \ i. \ k > j. \ x \ X2. zero x ([xs !! j]\<^sub>\) \ zero x ([xs !! k]\<^sub>\) \ (\ m. j < m \ m < k \ \ zero x ([xs !! m]\<^sub>\)) \ (\ x \ X2. \ m. j < m \ m \ k \ zero x ([xs !! m]\<^sub>\)) \ (\ x \ X1. \ m \ j. unbounded x ([xs !! m]\<^sub>\))" for i proof - from unbounded obtain i' where i': "\ x \ X1. \ m \ i'. unbounded x ([xs !! m]\<^sub>\)" by auto then obtain i' where i': "i' \ i" "\ x \ X1. \ m \ i'. unbounded x ([xs !! m]\<^sub>\)" by (cases "i' \ i"; auto) from finite(1) have "finite X2" unfolding X2_def by auto with \X2 \ {}\ \_divergent_divergent_aux2[where M = "{\ i. zero x ([xs !! i]\<^sub>\) | x. x \ X2}"] inf_resets have "\j\i'. \k>j. \P\{\i. zero x ([xs !! i]\<^sub>\) |x. x \ X2}. P j \ P k \ (\m \ P m) \ (\Q\{\i. zero x ([xs !! i]\<^sub>\) |x. x \ X2}. \m\k. j < m \ Q m)" by force then obtain j k x where "j \ i'" "k > j" "x \ X2" "zero x ([xs !! j]\<^sub>\)" "zero x ([xs !! k]\<^sub>\)" "\m. j < m \ m < k \ \ zero x ([xs !! m]\<^sub>\)" "\Q\{\i. zero x ([xs !! i]\<^sub>\) |x. x \ X2}. \m\k. j < m \ Q m" by auto moreover from this(7) have "\x\X2. \m \ k. j < m \ zero x ([xs !! m]\<^sub>\)" by auto ultimately show ?thesis using i' by (inst_existentials j k x) auto qed moreover have "\ j' \ j. dur xs j' - dur xs i \ 0.5" if x: "x \ X2" "i < j" "zero x ([xs !! i]\<^sub>\)" "zero x ([xs !! j]\<^sub>\)" and not_reset: "\ m. i < m \ m < j \ \ zero x ([xs !! m]\<^sub>\)" and X2: "\ x \ X2. \ m. i < m \ m \ j \ zero x ([xs !! m]\<^sub>\)" and X1: "\ x \ X1. \ m \ i. unbounded x ([xs !! m]\<^sub>\)" for x i j proof - have "\j'>j. \ zero x ([xs !! j']\<^sub>\)" proof - from inf_resets[OF x(1)] obtain j' where "j' \ Suc j" "\ zero x ([xs !! j']\<^sub>\)" by auto then show ?thesis by - (rule exI[where x = j']; auto) qed from inf_resets[OF x(1)] obtain j' where "j' \ Suc j" "\ zero x ([xs !! j']\<^sub>\)" by auto with nat_eventually_critical_path[OF x(4) this(2)] obtain j' where j': "j' > j" "\ zero x ([xs !! j']\<^sub>\)" "\ m \ j. m < j' \ zero x ([xs !! m]\<^sub>\)" by auto from \x \ X2\ have "x \ \" unfolding X2_def by simp with \i < j\ not_reset not_reset_dur \stream_trans _\ in_S pairwise_Suc[OF eq_elapsed] have "dur xs (j - 1) - dur xs i = (xs !! (j - 1)) x - (xs !! i) x" (is "?d1 = ?d2") by (auto simp: stream.pred_set) moreover from \zero x ([xs !! i]\<^sub>\)\ in_S have "(xs !! i) x = 0" by (auto intro: zeroD simp: stream.pred_set) ultimately have "dur xs (j - 1) - dur xs i = (xs !! (j - 1)) x" (is "?d1 = ?d2") by simp show ?thesis proof (cases "?d1 \ 0.5") case True (* XXX Fix SMT *) with dur_mono[of "j - 1" j xs] have "5 / 10 \ dur xs j - dur xs i" by simp then show ?thesis by blast next case False have j_c_bound: "(xs !! j) c \ ?d2" if "c \ X2" for c proof (cases "(xs !! j) c = 0") case True from in_S \j > _\ True \x \ \\ show ?thesis by (auto simp: V_def stream.pred_set) next case False from X2 \c \ X2\ in_S have "\k>i. k \ j \ (xs !! k) c = 0" by (force simp: zeroD stream.pred_set) with False have "\k>i. k \ j - Suc 0 \ (xs !! k) c = 0" by (metis Suc_le_eq Suc_pred linorder_neqE_nat not_less not_less_zero) moreover from that have "c \ \" by (auto simp: X2_def) moreover from not_reset in_S \x \ \\ have "\k>i. k \ j - 1 \ (xs !! k) x \ 0" by (auto simp: zeroI stream.pred_set) ultimately have "(xs !! (j - 1)) c \ ?d2" using trans in_S \_ x = 0\ \x \ \\ by (auto intro: \_divergent_divergent_aux that simp: stream.pred_set) moreover from trans_not_delay_mono[OF pairwise_Suc[OF trans], of "j - 1"] \x \ \\ \c \ \\ \j > _\ in_S x(4) have "(xs !! j) c \ (xs !! (j - 1)) c" by (auto simp: zeroD stream.pred_set) ultimately show ?thesis by auto qed moreover from False \?d1 = ?d2\ have "?d2 < 1" by auto moreover from in_S have "(xs !! j) c \ 0" if "c \ \" for c using that by (auto simp: V_def stream.pred_set) ultimately have frac_bound: "frac ((xs !! j) c) \ ?d2" if "c \ X2" for c using that frac_le_1I by (force simp: X2_def) let ?u = "(xs !! j)" from in_S have "[xs !! j]\<^sub>\ \ \" by (auto simp: stream.pred_set) then obtain I r where region: "[xs !! j]\<^sub>\ = region \ I r" "valid_region \ k I r" unfolding \_def by auto let ?S = "{frac (?u c) | c. c \ \ \ isIntv (I c)}" have \_X2: "c \ X2" if "c \ \" "isIntv (I c)" for c proof - from X1 \j > i\ have "\x\X1. unbounded x ([xs !! j]\<^sub>\)" by auto with unbounded_Greater[OF region(2) \c \ \\] region(1) that(2) have "c \ X1" by auto with \c \ \\ show "c \ X2" unfolding X2_def by auto qed have frac_bound: "frac ((xs !! j) c) \ ?d2" if "c \ \" "isIntv (I c)" for c using frac_bound[OF \_X2] that . have "dur xs (j' - 1) = dur xs j" using j' \x \ \\ in_S eq_elapsed by (subst dur_zero_tail[where \ = "smap (\ u. [u]\<^sub>\) xs"]) (auto dest: pairwise_Suc simp: stream.pred_set) moreover from dur_reset[OF eq_elapsed in_S, of x "j - 1"] \x \ \\ x(4) \j > _\ have "dur xs j = dur xs (j - 1)" by (auto simp: stream.pred_set) ultimately have "dur xs (j' - 1) = dur xs (j - 1)" by auto moreover have "dur xs j' - dur xs (j' - 1) \ (1 - ?d2) / 2" proof - from \j' > _\ have "j' > 0" by auto with pairwise_Suc[OF trans', of "j' - 1"] have "trans' (xs !! (j' - 1)) (xs !! j')" by auto moreover from j' have "(xs !! (j' - 1)) x = 0" "(xs !! j') x > 0" using in_S \x \ \\ by (force intro: zeroD dest: not_zeroD simp: stream.pred_set)+ moreover note delayedR_aux = calculation obtain t where "(xs !! j') = (xs !! (j' - 1)) \ t" "t \ (1 - ?d2) / 2" "t \ 0" proof - from in_S have "[xs !! j']\<^sub>\ \ \" by (auto simp: stream.pred_set) then obtain I' r' where region': "[xs !! j']\<^sub>\ = region \ I' r'" "valid_region \ k I' r'" unfolding \_def by auto let ?S' = "{frac ((xs !! (j' - 1)) c) |c. c \ \ \ Regions.isIntv (I' c)}" from finite(1) have "?d2 \ Max (?S' \ {0})" apply - apply (rule Max.boundedI) apply fastforce apply fastforce apply safe subgoal premises prems for _ c d proof - from j' have "(xs !! (j' - 1)) c = ?u c \ (xs !! (j' - 1)) c = 0" by (intro resets_mono'[OF eq_elapsed in_S trans _ \x \ \\ \c \ \\]; auto) then show ?thesis proof (standard, goal_cases) case A: 1 show ?thesis proof (cases "c \ X1") case True with X1 \j' > j\ \j > i\ have "unbounded c ([xs !! j']\<^sub>\)" by auto with region' \c \ \\ have "I' c = Greater (k c)" by (auto intro: unbounded_Greater) with prems show ?thesis by auto next case False with \c \ \\ have "c \ X2" unfolding X2_def by auto with j_c_bound have mono: "(xs !! j) c \ (xs !! (j - 1)) x" . from in_S \c \ \\ have "(xs !! (j' - 1)) c \ 0" unfolding V_def stream.pred_set by auto then have "frac ((xs !! (j' - 1)) c) \ (xs !! (j' - 1)) c" using frac_le_self by auto with A mono show ?thesis by auto qed next case prems: 2 (* XXX *) have "frac (0 :: real) = (0 :: real)" by auto then have "frac (0 :: real) \ (0 :: real)" by linarith moreover from in_S \x \ \\ have "(xs !! (j - 1)) x \ 0" unfolding V_def stream.pred_set by auto ultimately show ?thesis using prems by auto qed qed using in_S \x \ \\ by (auto simp: V_def stream.pred_set) then have le: "(1 - ?d2) / 2 \ (1 - Max (?S' \ {0})) / 2" by simp let ?u = "xs !! j'" let ?u' = "xs !! (j' - 1)" from in_S have *: "?u' \ V" "[?u']\<^sub>\ \ \" "?u \ V" "[?u]\<^sub>\ \ \" by (auto simp: stream.pred_set) from pairwise_Suc[OF trans, of "j' - 1"] \j' > j\ have "trans (xs !! (j' - 1)) (xs !! j')" by auto then have Succ: "[xs !! j']\<^sub>\ \ Succ \ ([xs !! (j' - 1)]\<^sub>\) \ (\ t\ 0. ?u = ?u' \ t)" proof cases case prems: (succ t) from * have "?u' \ [?u']\<^sub>\" by auto with prems * show ?thesis by auto next case (reset l) with \?u' \ V\ have "?u x \ ?u' x" by (cases "x \ set l") (auto simp: V_def) from j' have "zero x ([?u']\<^sub>\)" by auto with \?u' \ V\ have "?u' x = 0" unfolding zero_def by auto with \?u x \ _\ \?u x > 0\ show ?thesis by auto next case id with * Succ_refl[of \ \ k, folded \_def, OF _ finite(1)] show ?thesis unfolding cval_add_def by auto qed then obtain t where t: "?u = xs !! (j' - 1) \ t" "t \ 0" by auto note Succ = Succ[THEN conjunct1] show ?thesis proof (cases "\ c \ X2. \ d :: nat. ?u c = d") case True from True obtain c and d :: nat where c: "c \ \" "c \ X2" "?u c = d" by (auto simp: X2_def) have "?u x > 0" by fact from pairwise_Suc[OF eq_elapsed, of "j' - 1"] \j' > j\ have "eq_elapsed (xs !! (j' - 1)) ?u" by auto moreover from elapsed_eq[OF this \x \ \\] \(xs !! (j' - 1)) x = 0\ \(xs !! j') x > 0\ have "elapsed (xs !! (j' - 1)) (xs !! j') > 0" by auto ultimately have "?u c - (xs !! (j' - 1)) c > 0" using \c \ \\ unfolding eq_elapsed_def by auto moreover from in_S have "xs !! (j' - 1) \ V" by (auto simp: stream.pred_set) ultimately have "?u c > 0" using \c \ \\ unfolding V_def by auto from region' in_S \c \ \\ have "intv_elem c ?u (I' c)" by (force simp: stream.pred_set) with \?u c = d\ \?u c > 0\ have "?u c \ 1" by auto moreover have "(xs !! (j' - 1)) c \ 0.5" proof - have "(xs !! (j' - 1)) c \ (xs !! j) c" (* XXX This proof is duplicated at least once above *) using j'(1,3) by (auto intro: resets_mono[OF eq_elapsed in_S trans _ \x \ \\ \c \ \\]) also have "\ \ ?d2" using j_c_bound[OF \c \ X2\] . also from \?d1 = ?d2\ \\ 5 / 10 \ _\ have "\ \ 0.5" by simp finally show ?thesis . qed moreover have "?d2 \ 0" using in_S \x \ \\ by (auto simp: V_def stream.pred_set) ultimately have "?u c - (xs !! (j' - 1)) c \ (1 - ?d2) / 2" by auto with t have "t \ (1 - ?d2) / 2" unfolding cval_add_def by auto with t show ?thesis by (auto intro: that) next case F: False have not_const: "\ isConst (I' c)" if "c \ \" for c proof (rule ccontr, simp) assume A: "isConst (I' c)" show False proof (cases "c \ X1") case True with X1 \j' > j\ \j > _\ have "unbounded c ([xs !! j']\<^sub>\)" by auto with unbounded_Greater \c \ \\ region' have "isGreater (I' c)" by force with A show False by auto next case False with \c \ \\ have "c \ X2" unfolding X2_def by auto from region' in_S \c \ \\ have "intv_elem c ?u (I' c)" unfolding stream.pred_set by force with \c \ X2\ A False F show False by auto qed qed have "\x. x \ k c \ (xs !! j') c = real x" if "c \ \" for c proof (cases "c \ X2"; safe) fix d assume "c \ X2" "(xs !! j') c = real d" with F show False by auto next fix d assume "c \ X2" with that have "c \ X1" unfolding X2_def by auto with X1 \j' > j\ \j > i\ have "unbounded c ([?u]\<^sub>\)" by auto from unbounded_all[OF _ _ this] \c \ \\ in_S have "?u c > k c" by (force simp: stream.pred_set) moreover assume "?u c = real d" "d \ k c" ultimately show False by auto qed with delayedR_aux have "(xs !! j') = delayedR ([xs !! j']\<^sub>\) (xs !! (j' - 1))" using \x \ \\ unfolding trans'_def by auto from not_const region'(1) in_S Succ(1) have "\t\0. delayedR ([xs !! j']\<^sub>\) (xs !! (j' - 1)) = xs !! (j' - 1) \ t \ (1 - Max (?S' \ {0})) / 2 \ t" apply simp apply (rule delayedR_correct(2)[OF _ _ region'(2), simplified]) by (auto simp: stream.pred_set) with le \_ = delayedR _ _\ show ?thesis by (auto intro: that) qed qed moreover from pairwise_Suc[OF eq_elapsed, of "j' - 1"] \j' > 0\ have "eq_elapsed (xs !! (j' - 1)) (xs !! j')" by auto ultimately show "dur xs j' - dur xs (j' - 1) \ (1 - ?d2) / 2" using \j' > 0\ dur_Suc[of _ "j' - 1"] \x \ \\ by (auto simp: cval_add_def elapsed_eq) qed moreover from dur_mono[of i "j - 1" xs] \i < j\ have "dur xs i \ dur xs (j - 1)" by simp ultimately have "dur xs j' - dur xs i \ 0.5" unfolding \?d1 = ?d2\[symmetric] by auto then show ?thesis using \j < j'\ by - (rule exI[where x = j']; auto) qed qed moreover have "\ j' \ i. dur xs j' - dur xs i \ 0.5" for i proof - from calculation(1)[of i] obtain j k x where "j\i" "k>j" "x\X2" "zero x ([xs !! j]\<^sub>\)" "zero x ([xs !! k]\<^sub>\)" "\m. j < m \ m < k \ \ zero x ([xs !! m]\<^sub>\)" "\x\X2. \m>j. m \ k \ zero x ([xs !! m]\<^sub>\)" "\x\X1. \m\j. unbounded x ([xs !! m]\<^sub>\)" by auto from calculation(2)[OF this(3,2,4-8)] obtain j' where "j'\k" "5 / 10 \ dur xs j' - dur xs j" by auto with dur_mono[of i j xs] \j \ i\ \k > j\ show ?thesis by (intro exI[where x = j']; auto) qed then show ?thesis by - (rule dur_ev_exceedsI[where d = "0.5"]; auto) qed qed lemma cfg_on_div_absc: notes in_space_UNIV[measurable] assumes "cfg \ cfg_on_div st" "st \ S" shows "absc cfg \ R_G_cfg_on_div (abss st)" proof - from assms have *: "cfg \ MDP.cfg_on st" "state cfg = st" "div_cfg cfg" unfolding cfg_on_div_def by auto with assms have "cfg \ valid_cfg" by (auto intro: MDP.valid_cfgI) have "almost_everywhere (MDP.MC.T cfg) (MDP.MC.enabled cfg)" by (rule MDP.MC.AE_T_enabled) moreover from * have "AE x in MDP.MC.T cfg. divergent (smap (snd \ state) x)" by (simp add: div_cfg_def) ultimately have "AE x in MDP.MC.T cfg. \_div (smap (snd \ state) (smap absc x))" proof eventually_elim case (elim \) let ?xs = "smap (snd o state) \" from MDP.pred_stream_cfg_on[OF \_ \ valid_cfg\ \MDP.MC.enabled _ _\] have *: "pred_stream (\ x. x \ S) (smap state \)" by (auto simp: stream.pred_set) have "[snd (state x)]\<^sub>\ = snd (abss (state x))" if "x \ sset \" for x proof - from * that have "state x \ S" by (auto simp: stream.pred_set) then have "snd (abss (state x)) = [snd (state x)]\<^sub>\" by (metis abss_S snd_conv surj_pair) then show ?thesis .. qed then have "smap (\z. [snd (state z)]\<^sub>\) \ = (smap (\z. snd (abss (state z))) \)" by auto from * have "pred_stream (\ u. u \ V) ?xs" apply (simp add: map_def stream.pred_set) apply (subst (asm) surjective_pairing) using S_V by blast moreover have "stream_trans ?xs" by (rule enabled_stream_trans \_ \ valid_cfg\ \MDP.MC.enabled _ _\)+ ultimately show ?case using \divergent _\ \smap _ \ = _\ by - (drule divergent_\_divergent, auto simp add: stream.map_comp state_absc) qed with \cfg \ valid_cfg\ have "R_G_div_cfg (absc cfg)" unfolding R_G_div_cfg_def by (subst absc_distr_self) (auto intro: MDP.valid_cfgI simp: AE_distr_iff) with R_G.valid_cfgD \cfg \ valid_cfg\ * show ?thesis unfolding R_G_cfg_on_div_def by auto force qed definition "alternating cfg = (AE \ in MDP.MC.T cfg. alw (ev (HLD {cfg. \cfg' \ K_cfg cfg. fst (state cfg') = fst (state cfg)})) \)" lemma K_cfg_same_loc_iff: "(\cfg'\ K_cfg cfg. fst (state cfg') = fst (state cfg)) \ (\cfg'\ K_cfg (absc cfg). fst (state cfg') = fst (state (absc cfg)))" if "cfg \ valid_cfg" using that by (auto simp: state_absc fst_abss K_cfg_map_absc) lemma (in -) stream_all2_flip: "stream_all2 (\a b. R b a) xs ys = stream_all2 R ys xs" by (standard; coinduction arbitrary: xs ys; auto dest: sym) (* TODO: Lots of duplication here, e.g. with path_measure_eq_repcs1_new *) lemma AE_alw_ev_same_loc_iff: assumes "cfg \ valid_cfg" shows "alternating cfg \ alternating (absc cfg)" unfolding alternating_def apply (simp add: MDP.MC.T.AE_iff_emeasure_eq_1) subgoal proof - show ?thesis (is "(?x = 1) = (?y = 1)") proof - have *: "stream_all2 (\s t. t = absc s) x y = stream_all2 (=) y (smap absc x)" for x y by (subst stream_all2_flip) simp have "?x = ?y" apply (rule T_eq_rel_half[where f = absc and S = valid_cfg, OF HOL.refl, rotated 2]) subgoal apply (simp add: space_stream_space rel_set_strong_def) apply (intro allI impI) apply (frule stream.rel_mono_strong[where Ra = "\s t. t = absc s"]) by (auto simp: * stream.rel_eq stream_all2_refl alw_holds_pred_stream_iff[symmetric] K_cfg_same_loc_iff HLD_def elim!: alw_ev_cong) subgoal by (rule rel_funI) (auto intro!: rel_pmf_reflI simp: pmf.rel_map(2) K_cfg_map_absc) using \cfg \ valid_cfg\ by simp+ then show ?thesis by simp qed qed done lemma AE_alw_ev_same_loc_iff': assumes "cfg \ R_G.cfg_on (abss st)" "st \ S" shows "alternating cfg \ alternating (repcs st cfg)" proof - from assms have "cfg \ R_G.valid_cfg" by (auto intro: R_G.valid_cfgI) with assms show ?thesis by (subst AE_alw_ev_same_loc_iff) (auto simp: absc_repcs_id) qed lemma (in -) cval_add_non_id: False if "b \ d = b" "d > 0" for d :: real proof - from that(1) have "(b \ d) x = b x" by (rule fun_cong) with \d > 0\ show False unfolding cval_add_def by simp qed lemma repcs_unbounded_AE_non_loop_end_strong: assumes "cfg \ R_G.cfg_on (abss st)" "st \ S" and "alternating cfg" shows "AE \ in MDP.MC.T (repcs st cfg). (\ u :: ('c \ real). (\ c \ \. u c > real (k c)) \ \ (ev (alw (\ xs. shd xs = u))) (smap (snd o state) \))" (is "AE \ in ?M. ?P \") proof - from assms have "cfg \ R_G.valid_cfg" by (auto intro: R_G.valid_cfgI) with assms(1) have "repcs st cfg \ valid_cfg" by auto from R_G.valid_cfgD[OF \cfg \ R_G.valid_cfg\] have "cfg \ R_G.cfg_on (state cfg)" . let ?U = "\ u. \ l \ L. {\ \ K (l, u). \ \ return_pmf (l, u) \ (\ x \ \. fst x = l)}" let ?r = "\ u. Sup ({0} \ (\ \. measure_pmf \ {x. snd x = u}) ` ?U u)" have lt_1: "?r u < 1" for u proof - have *: "emeasure (measure_pmf \) {x. snd x = u} < 1" if "\ \ return_pmf (l, u)" "\x\set_pmf \. fst x = l" for \ and l :: 's proof (rule ccontr) assume "\ emeasure (measure_pmf \) {x. snd x = u} < 1" then have "1 = emeasure (measure_pmf \) {x. snd x = u}" using measure_pmf.emeasure_ge_1_iff by force also from that(2) have "\ \ emeasure (measure_pmf \) {(l, u)}" by (subst emeasure_Int_set_pmf[symmetric]) (auto intro!: emeasure_mono) finally show False by (simp add: measure_pmf.emeasure_ge_1_iff measure_pmf_eq_1_iff that(1)) qed let ?S = "{map_pmf (\ (X, l). (l, ([X := 0]u))) \ | \ l g. (l, g, \) \ trans_of A}" have "(\ \. measure_pmf \ {x. snd x = u}) ` ?U u \ {0, 1} \ (\ \. measure_pmf \ {x. snd x = u}) ` ?S" by (force elim!: K.cases) moreover have "finite ?S" proof - have "?S \ (\ (l, g, \). map_pmf (\ (X, l). (l, ([X := 0]u))) \) ` trans_of A" by force also from finite(3) have "finite \" .. finally show ?thesis . qed ultimately have "finite ((\ \. measure_pmf \ {x. snd x = u}) ` ?U u)" by (auto intro: finite_subset) then show ?thesis by (fastforce intro: * finite_imp_Sup_less) qed { fix l :: 's and u :: "'c \ real" and cfg :: "('s \ ('c \ real) set) cfg" assume unbounded: "\ c \ \. u c > k c" and "cfg \ R_G.cfg_on (abss (l, u))" "abss (l, u) \ \" and same_loc: "\ cfg' \ K_cfg cfg. fst (state cfg') = l" then have "cfg \ R_G.valid_cfg" "repcs (l, u) cfg \ valid_cfg" by (auto intro: R_G.valid_cfgI) then have cfg_on: "repcs (l, u) cfg \ MDP.cfg_on (l, u)" by (auto dest: MDP.valid_cfgD) from \cfg \ R_G.cfg_on _\ have "action cfg \ \ (abss (l, u))" by (rule R_G.cfg_onD_action) (* TODO: Pull out? *) have K_cfg_rept: "state ` K_cfg (repcs (l, u) cfg) = rept (l, u) (action cfg)" unfolding K_cfg_def by (force simp: action_repcs) have "l \ L" using MDP.valid_cfg_state_in_S \repcs (l, u) cfg \ MDP.valid_cfg\ by fastforce moreover have "rept (l, u) (action cfg) \ return_pmf (l, u)" proof (rule ccontr, simp) assume "rept (l, u) (action cfg) = return_pmf (l, u)" then have "action cfg = return_pmf (abss (l, u))" using abst_rept_id[OF \action cfg \ _\] by (simp add: abst_def) moreover have "(l, u) \ S" using \_ \ \\ by (auto dest: \_abss_S) moreover have "abss (l, u) = (l, [u]\<^sub>\)" by (metis abss_S calculation(2)) ultimately show False using \rept (l, u) _ = _\ unbounded unfolding rept_def by (auto dest: cval_add_non_id) qed moreover have "rept (l, u) (action cfg) \ K (l, u)" proof - have "action (repcs (l, u) cfg) \ K (l, u)" using cfg_on by blast then show ?thesis by (simp add: repcs_def) qed moreover have "\x\set_pmf (rept (l, u) (action cfg)). fst x = l" using same_loc K_cfg_same_loc_iff[of "repcs (l, u) cfg"] \repcs (l, u) _ \ valid_cfg\ \cfg \ R_G.valid_cfg\ \cfg \ R_G.cfg_on _\ by (simp add: absc_repcs_id fst_abss K_cfg_rept[symmetric]) ultimately have "rept (l, u) (action cfg) \ ?U u" by blast then have "measure_pmf (rept (l, u) (action cfg)) {x. snd x = u} \ ?r u" by (fastforce intro: Sup_upper) moreover have "rept (l, u) (action cfg) = action (repcs (l, u) cfg)" by (simp add: repcs_def) ultimately have "measure_pmf (action (repcs (l, u) cfg)) {x. snd x = u} \ ?r u" by auto } note * = this let ?S = "{cfg. \ cfg' s. cfg' \ R_G.valid_cfg \ cfg = repcs s cfg' \ abss s = state cfg'}" have start: "repcs st cfg \ ?S" using \cfg \ R_G.valid_cfg\ assms unfolding R_G_cfg_on_div_def by clarsimp (inst_existentials cfg "fst st" "snd st", auto) have step: "y \ ?S" if "y \ K_cfg x" "x \ ?S" for x y using that apply safe subgoal for cfg' l u apply (inst_existentials "absc y" "state y") subgoal by blast subgoal by (metis K_cfg_valid_cfgD R_G.valid_cfgD R_G.valid_cfg_state_in_S absc_repcs_id cont_absc_1 cont_repcs1 repcs_valid ) subgoal by (simp add: state_absc) done done have **: "x \ ?S" if "(repcs st cfg, x) \ MDP.MC.acc" for x proof - from MDP.MC.acc_relfunD[OF that] obtain n where "((\ a b. b \ K_cfg a) ^^ n) (repcs st cfg) x" . then show ?thesis proof (induction n arbitrary: x) (* XXX Extract induction rule *) case 0 with start show ?case by simp next case (Suc n) from this(2)[simplified] show ?case apply (rule relcomppE) apply (erule step) apply (erule Suc.IH) done qed qed have ***: "almost_everywhere (MDP.MC.T (repcs st cfg)) (alw (HLD ?S))" by (rule AE_mp[OF MDP.MC.AE_T_reachable]) (fastforce dest: ** simp: HLD_iff elim: alw_mono) from \alternating cfg\ assms have "alternating (repcs st cfg)" by (simp add: AE_alw_ev_same_loc_iff'[of _ st]) then have alw_ev_same2: "almost_everywhere (MDP.MC.T (repcs st cfg)) (alw (\\. HLD (state -` snd -` {u}) \ \ ev (HLD {cfg. \cfg'\set_pmf (K_cfg cfg). fst (state cfg') = fst (state cfg)}) \))" for u unfolding alternating_def by (auto elim: alw_mono) let ?X = "{cfg :: ('s \ ('c \ real)) cfg. \ c \ \. snd (state cfg) c > k c}" let ?Y = "{cfg. \ cfg' \ K_cfg cfg. fst (state cfg') = fst (state cfg)}" have "(AE \ in ?M. ?P \) \ (AE \ in ?M. \ u :: ('c \ real). (\ c \ \. u c > k c) \ u \ snd ` state ` (MDP.MC.acc `` {repcs st cfg}) \ \ (ev (alw (\ xs. shd xs = u))) (smap (snd o state) \))" (is "?L \ ?R") proof assume ?L then show ?R by eventually_elim auto next assume ?R with MDP.MC.AE_T_reachable[of "repcs st cfg"] show ?L proof (eventually_elim, intro allI impI notI, goal_cases) case (1 \ u) then show ?case by - (intro alw_HLD_smap alw_disjoint_ccontr[where S = "(snd o state) ` MDP.MC.acc `` {repcs st cfg}" and R = "{u}" and \ = "smap (snd o state) \" ]; auto simp: HLD_iff) qed qed also have "\ \ (\ u :: ('c \ real). (\ c \ \. u c > k c) \ u \ snd ` state ` (MDP.MC.acc `` {repcs st cfg}) \ (AE \ in ?M. \ (ev (alw (\ xs. shd xs = u))) (smap (snd o state) \)))" using MDP.MC.countable_reachable[of "repcs st cfg"] by - (rule AE_all_imp_countable, auto intro: countable_subset[where B = "snd ` state ` MDP.MC.acc `` {repcs st cfg}"]) also show ?thesis unfolding calculation apply clarsimp subgoal for l u x apply (rule MDP.non_loop_tail_strong[simplified, of snd "snd (state x)" ?Y ?S "?r (snd (state x))"] ) subgoal apply safe subgoal premises prems for cfg l1 u1 _ cfg' l2 u2 proof - have [simp]: "l2 = l1" "u2 = u1" subgoal by (metis MDP.cfg_onD_state Pair_inject prems(4) state_repcs) subgoal by (metis MDP.cfg_onD_state prems(4) snd_conv state_repcs) done with prems have [simp]: "u2 = u" by (metis \(l, u) = state x\ \snd (l1, u1) = snd (state x)\ \u2 = u1\ snd_conv) have [simp]: "snd -` {snd (state x)} = {y. snd y = snd (state x)}" by (simp add: vimage_def) from prems show ?thesis apply simp apply (erule *[simplified]) subgoal using prems(1) prems(2)[symmetric] prems(3-) by (auto simp: R_G.valid_cfg_def) subgoal using prems(1) prems(2)[symmetric] prems(3-) by (auto simp: R_G.valid_cfg_def) subgoal using K_cfg_same_loc_iff[of "repcs (l1, snd (state x)) cfg'"] by (simp add: absc_repcs_id) (metis fst_abss fst_conv repcs_valid) done qed done subgoal by (auto intro: lt_1[simplified]) apply (rule MDP.valid_cfgD[OF \repcs st cfg \ valid_cfg\]; fail) subgoal using *** unfolding alw_holds_pred_stream_iff[symmetric] HLD_def . subgoal by (rule alw_ev_same2) done done qed lemma cfg_on_div_repcs_strong: notes in_space_UNIV[measurable] assumes "cfg \ R_G_cfg_on_div (abss st)" "st \ S" and "alternating cfg" shows "repcs st cfg \ cfg_on_div st" proof - let ?st = "abss st" let ?cfg = "repcs st cfg" from assms have *: "cfg \ R_G.cfg_on ?st" "state cfg = ?st" "R_G_div_cfg cfg" unfolding R_G_cfg_on_div_def by auto with assms have "cfg \ R_G.valid_cfg" by (auto intro: R_G.valid_cfgI) with \st \ S\ \_ = ?st\ have "?cfg \ valid_cfg" by auto from *(1) \st \ S\ \alternating cfg\ have "AE \ in MDP.MC.T ?cfg. \u. (\c\\. real (k c) < u c) \ \ ev (alw (\xs. shd xs = u)) (smap (snd \ state) \)" by (rule repcs_unbounded_AE_non_loop_end_strong) \ \Move to lower level\ moreover from *(2,3) have "AE \ in MDP.MC.T ?cfg. \_div (smap (snd \ state) (smap absc \))" unfolding R_G_div_cfg_def by (subst (asm) R_G_trace_space_distr_eq[OF \cfg \ R_G.valid_cfg\]; simp add: AE_distr_iff) ultimately have "div_cfg ?cfg" unfolding div_cfg_def using MDP.MC.AE_T_enabled[of ?cfg] proof eventually_elim case prems: (elim \) let ?xs = "smap (snd o state) \" from MDP.pred_stream_cfg_on[OF \_ \ valid_cfg\ \MDP.MC.enabled _ _\] have *: "pred_stream (\ x. x \ S) (smap state \)" by (auto simp: stream.pred_set) have "[snd (state x)]\<^sub>\ = snd (abss (state x))" if "x \ sset \" for x proof - from * that have "state x \ S" by (auto simp: stream.pred_set) then have "snd (abss (state x)) = [snd (state x)]\<^sub>\" by (metis abss_S snd_conv surj_pair) then show ?thesis .. qed then have "smap (\z. [snd (state z)]\<^sub>\) \ = (smap (\z. snd (abss (state z))) \)" by auto from * have "pred_stream (\ u. u \ V) ?xs" by (simp add: map_def stream.pred_set, subst (asm) surjective_pairing, blast) moreover have "stream_trans ?xs" by (rule enabled_stream_trans \_ \ valid_cfg\ \MDP.MC.enabled _ _\)+ moreover have "pairwise trans' ?xs" using \_ \ R_G.valid_cfg\ \state cfg = _\[symmetric] \MDP.MC.enabled _ _\ by (rule enabled_stream_trans') moreover from prems(1) have "\u. (\c\\. real (k c) < u c) \ \ ev (alw (\xs. snd (shd xs) = u)) (smap state \)" by simp ultimately show ?case using \\_div _\ by (simp add: stream.map_comp state_absc \smap _ \ = _\ \_divergent_divergent) qed with MDP.valid_cfgD \cfg \ R_G.valid_cfg\ * show ?thesis unfolding cfg_on_div_def by auto force qed lemma repcs_unbounded_AE_non_loop_end: assumes "cfg \ R_G.cfg_on (abss st)" "st \ S" shows "AE \ in MDP.MC.T (repcs st cfg). (\ s :: ('s \ ('c \ real)). (\ c \ \. snd s c > k c) \ \ (ev (alw (\ xs. shd xs = s))) (smap state \))" (is "AE \ in ?M. ?P \") proof - from assms have "cfg \ R_G.valid_cfg" by (auto intro: R_G.valid_cfgI) with assms(1) have "repcs st cfg \ valid_cfg" by auto from R_G.valid_cfgD[OF \cfg \ R_G.valid_cfg\] have "cfg \ R_G.cfg_on (state cfg)" . let ?K = "\ x. {\ \ K x. \ \ return_pmf x}" let ?r = "\ x. Sup ((\ \. measure_pmf \ {x}) ` ?K x)" have lt_1: "?r x < 1" if "\ \ ?K x" for \ x proof - have *: "emeasure (measure_pmf \) {x} < 1" if "\ \ return_pmf x" for \ proof (rule ccontr) assume "\ emeasure (measure_pmf \) {x} < 1" then have "emeasure (measure_pmf \) {x} = 1" using measure_pmf.emeasure_ge_1_iff by force with that show False by (simp add: measure_pmf_eq_1_iff) qed let ?S = "{map_pmf (\ (X, l). (l, ([X := 0]u))) \ | \ l u g. x = (l, u) \ (l, g, \) \ trans_of A}" have "(\ \. measure_pmf \ {x}) ` ?K x \ {0, 1} \ (\ \. measure_pmf \ {x}) ` ?S" by (force elim!: K.cases) moreover have "finite ?S" proof - have "?S \ (\ (l, g, \). map_pmf (\ (X, l). (l, ([X := 0](snd x)))) \) ` trans_of A" by force also from finite(3) have "finite \" .. finally show ?thesis . qed ultimately have "finite ((\ \. measure_pmf \ {x}) ` ?K x)" by (auto intro: finite_subset) then show ?thesis using that by (auto intro: * finite_imp_Sup_less) qed { fix s :: "'s \ ('c \ real)" and cfg :: "('s \ ('c \ real) set) cfg" assume unbounded: "\ c \ \. snd s c > k c" and "cfg \ R_G.cfg_on (abss s)" "abss s \ \" then have "repcs s cfg \ valid_cfg" by (auto intro: R_G.valid_cfgI) then have cfg_on: "repcs s cfg \ MDP.cfg_on s" by (auto dest: MDP.valid_cfgD) from \cfg \ _\ have "action cfg \ \ (abss s)" by (rule R_G.cfg_onD_action) have "rept s (action cfg) \ return_pmf s" proof (rule ccontr, simp) assume "rept s (action cfg) = return_pmf s" then have "action cfg = return_pmf (abss s)" using abst_rept_id[OF \action cfg \ _\] by (simp add: abst_def) moreover have "(fst s, snd s) \ S" using \_ \ \\ by (auto dest: \_abss_S) moreover have "abss s = (fst s, [snd s]\<^sub>\)" by (metis abss_S calculation(2) prod.collapse) ultimately show False using \rept s _ = _\ unbounded unfolding rept_def by (cases s) (auto dest: cval_add_non_id) qed moreover have "rept s (action cfg) \ K s" proof - have "action (repcs s cfg) \ K s" using cfg_on by blast then show ?thesis by (simp add: repcs_def) qed ultimately have "rept s (action cfg) \ ?K s" by blast then have "measure_pmf (rept s (action cfg)) {s} \ ?r s" by (auto intro: Sup_upper) moreover have "rept s (action cfg) = action (repcs s cfg)" by (simp add: repcs_def) ultimately have "measure_pmf (action (repcs s cfg)) {s} \ ?r s" by auto note this \rept s (action cfg) \ ?K s\ } note * = this let ?S = "{cfg. \ cfg' s. cfg' \ R_G.valid_cfg \ cfg = repcs s cfg' \ abss s = state cfg'}" have start: "repcs st cfg \ ?S" using \cfg \ R_G.valid_cfg\ assms unfolding R_G_cfg_on_div_def by clarsimp (inst_existentials cfg "fst st" "snd st", auto) have step: "y \ ?S" if "y \ K_cfg x" "x \ ?S" for x y using that apply safe subgoal for cfg' l u apply (inst_existentials "absc y" "state y") subgoal by blast subgoal by (metis K_cfg_valid_cfgD R_G.valid_cfgD R_G.valid_cfg_state_in_S absc_repcs_id cont_absc_1 cont_repcs1 repcs_valid ) subgoal by (simp add: state_absc) done done have **: "x \ ?S" if "(repcs st cfg, x) \ MDP.MC.acc" for x proof - from MDP.MC.acc_relfunD[OF that] obtain n where "((\ a b. b \ K_cfg a) ^^ n) (repcs st cfg) x" . then show ?thesis proof (induction n arbitrary: x) (* XXX Extract induction rule *) case 0 with start show ?case by simp next case (Suc n) from this(2)[simplified] show ?case by (elim relcomppE step Suc.IH) qed qed have ***: "almost_everywhere (MDP.MC.T (repcs st cfg)) (alw (HLD ?S))" by (rule AE_mp[OF MDP.MC.AE_T_reachable]) (fastforce dest: ** simp: HLD_iff elim: alw_mono) have "(AE \ in ?M. ?P \) \ (AE \ in ?M. \ s :: ('s \ ('c \ real)). (\ c \ \. snd s c > k c) \ s \ state ` (MDP.MC.acc `` {repcs st cfg}) \ \ (ev (alw (\ xs. shd xs = s))) (smap state \))" (is "?L \ ?R") proof assume ?L then show ?R by eventually_elim auto next assume ?R with MDP.MC.AE_T_reachable[of "repcs st cfg"] show ?L proof (eventually_elim, intro allI impI notI, goal_cases) case (1 \ s) from this(1,2,5,6) show ?case by (intro alw_HLD_smap alw_disjoint_ccontr[where S = "state ` MDP.MC.acc `` {repcs st cfg}" and R = "{s}" and \ = "smap state \" ]; simp add: HLD_iff; blast) qed qed also have "\ \ (\ s :: ('s \ ('c \ real)). (\ c \ \. snd s c > k c) \ s \ state ` (MDP.MC.acc `` {repcs st cfg}) \ (AE \ in ?M. \ (ev (alw (\ xs. shd xs = s))) (smap state \)))" using MDP.MC.countable_reachable[of "repcs st cfg"] by - (rule AE_all_imp_countable, auto intro: countable_subset[where B = "state ` MDP.MC.acc `` {repcs st cfg}"]) also show ?thesis unfolding calculation apply clarsimp subgoal for l u x apply (rule MDP.non_loop_tail'[simplified, of "state x" ?S "?r (state x)"]) subgoal apply safe subgoal premises prems for cfg cfg' l' u' proof - from prems have "state x = (l', u')" by (metis MDP.cfg_onD_state state_repcs) with \_ = state x\ have [simp]: "l = l'" "u = u'" by auto show ?thesis unfolding \state x = _\ using prems(1,3-) by (auto simp: R_G.valid_cfg_def intro: *) qed done subgoal apply (drule **) apply clarsimp apply (rule lt_1) apply (rule *) apply (auto dest: R_G.valid_cfg_state_in_S R_G.valid_cfgD) done apply (rule MDP.valid_cfgD[OF \repcs st cfg \ valid_cfg\]; fail) using *** unfolding alw_holds_pred_stream_iff[symmetric] HLD_def . done qed end (* PTA Regions *) subsection \Main Result\ text_raw \ \label{thm:minmax} \ context Probabilistic_Timed_Automaton_Regions_Reachability begin lemma R_G_cfg_on_valid: "cfg \ R_G.valid_cfg" if "cfg \ R_G_cfg_on_div s'" using that unfolding R_G_cfg_on_div_def R_G.valid_cfg_def by auto lemma cfg_on_valid: "cfg \ valid_cfg" if "cfg \ cfg_on_div s" using that unfolding cfg_on_div_def MDP.valid_cfg_def by auto abbreviation "path_measure P cfg \ emeasure (MDP.T cfg) {x\space MDP.St. P x}" abbreviation "R_G_path_measure P cfg \ emeasure (R_G.T cfg) {x\space R_G.St. P x}" abbreviation "progressive st \ cfg_on_div st \ {cfg. alternating cfg}" abbreviation "R_G_progressive st \ R_G_cfg_on_div st \ {cfg. alternating cfg}" text \Summary of our results on divergent configurations:\ lemma absc_valid_cfg_eq: "absc ` progressive s = R_G_progressive s'" apply safe subgoal unfolding s'_def by (rule cfg_on_div_absc) auto subgoal by (simp add: AE_alw_ev_same_loc_iff cfg_on_valid) subgoal for cfg unfolding s'_def by (frule cfg_on_div_repcs_strong) (auto 4 4 simp: s'_def R_G_cfg_on_div_def AE_alw_ev_same_loc_iff'[symmetric] intro: R_G_cfg_on_valid absc_repcs_id[symmetric] ) done text \Main theorem:\ theorem Min_Max_reachability: notes in_space_UNIV[measurable] and [iff] = pred_stream_iff shows "(\cfg\ progressive s. path_measure (\ x. (holds \ suntil holds \) (s ## x)) cfg) = (\cfg\ R_G_progressive s'. R_G_path_measure (\ x. (holds \' suntil holds \') (s' ## x)) cfg) \ (\cfg\ progressive s. path_measure (\ x. (holds \ suntil holds \) (s ## x)) cfg) = (\cfg\ R_G_progressive s'. R_G_path_measure (\ x. (holds \' suntil holds \') (s' ## x)) cfg)" proof (rule SUP_eq_and_INF_eq; rule bexI[rotated]; erule IntE) fix cfg assume cfg_div: "cfg \ R_G_cfg_on_div s'" and "cfg \ Collect alternating" then have "alternating cfg" by auto let ?cfg' = "repcs s cfg" from \alternating cfg\ cfg_div have "alternating ?cfg'" by (simp add: R_G_cfg_on_div_def s'_def AE_alw_ev_same_loc_iff'[of _ s]) with cfg_div \alternating cfg\ show "?cfg' \ cfg_on_div s \ Collect alternating" by (auto intro: cfg_on_div_repcs_strong simp: s'_def) show "emeasure (R_G.T cfg) {x \ space R_G.St. (holds \' suntil holds \') (s' ## x)} = emeasure (MDP.T ?cfg') {x \ space MDP.St. (holds \ suntil holds \) (s ## x)}" (is "?a = ?b") proof - from cfg_div have "cfg \ R_G.valid_cfg" by (rule R_G_cfg_on_valid) from cfg_div have "cfg \ R_G.cfg_on s'" unfolding R_G_cfg_on_div_def by auto then have "state cfg = s'" by auto have "?a = ?b" apply (rule path_measure_eq_repcs''_new[ of s cfg \ \, folded \'_def \'_def, unfolded \_ = s'\ state_repcs ] ) subgoal unfolding s'_def .. subgoal by fact subgoal using \?cfg' \ cfg_on_div s \ _\ by (blast intro: cfg_on_valid) subgoal premises prems for xs using prems s by (intro \_stream) subgoal premises prems using prems s by (intro \_stream) done then show ?thesis by simp qed next fix cfg assume cfg_div: "cfg \ cfg_on_div s" and "cfg \ Collect alternating" with absc_valid_cfg_eq show "absc cfg \ R_G_cfg_on_div s' \ Collect alternating" by auto show "emeasure (MDP.T cfg) {x \ space MDP.St. (holds \ suntil holds \) (s ## x)} = emeasure (R_G.T (absc cfg)) {x \ space R_G.St. (holds \' suntil holds \') (s' ## x)}" (is "?a = ?b") proof - have "absc cfg \ R_G.valid_cfg" using R_G_cfg_on_valid \absc cfg \ R_G_cfg_on_div s' \ _\ by blast from cfg_div have "cfg \ valid_cfg" by (simp add: cfg_on_valid) with \absc cfg \ R_G.valid_cfg\ have "?b = ?a" by (intro MDP.alw_S R_G.alw_S path_measure_eq_absc1_new [where P = "pred_stream (\s. s \ \)" and Q = "pred_stream (\s. s \ S)"] ) (auto simp: S_abss_\ intro: \_abss_S intro!: suntil_abss suntil_reps, measurable) then show "?a = ?b" by simp qed qed end (* PTA Reachability Problem *) end (* Theory *) diff --git a/thys/Probabilistic_Timed_Automata/library/Basic.thy b/thys/Probabilistic_Timed_Automata/library/Basic.thy --- a/thys/Probabilistic_Timed_Automata/library/Basic.thy +++ b/thys/Probabilistic_Timed_Automata/library/Basic.thy @@ -1,95 +1,95 @@ (* Author: Julian Brunner *) (* This is originally a part of the CAVA model checker *) theory Basic imports Main begin abbreviation (input) "const x \ \ _. x" lemma infinite_subset[trans]: "infinite A \ A \ B \ infinite B" using infinite_super by this lemma finite_subset[trans]: "A \ B \ finite B \ finite A" using finite_subset by this declare infinite_coinduct[case_names infinite, coinduct pred: infinite] lemma infinite_psubset_coinduct[case_names infinite, consumes 1]: assumes "R A" assumes "\ A. R A \ \ B \ A. R B" shows "infinite A" proof assume "finite A" then show "False" using assms by (induct rule: finite_psubset_induct) (auto) qed lemma GreatestI: fixes k :: nat assumes "P k" "\ k. P k \ k \ l" shows "P (GREATEST k. P k)" - proof (rule GreatestI_nat, safe) + proof (rule GreatestI_nat) show "P k" using assms(1) by this show "k \ l" if "P k" for k using assms(2) that by force qed lemma Greatest_le: fixes k :: nat assumes "P k" "\ k. P k \ k \ l" shows "k \ (GREATEST k. P k)" - proof (rule Greatest_le_nat, safe) + proof (rule Greatest_le_nat) show "P k" using assms(1) by this show "k \ l" if "P k" for k using assms(2) that by force qed lemma Greatest_not_less: fixes k :: nat assumes "k > (GREATEST k. P k)" "\ k. P k \ k \ l" shows "\ P k" proof assume 1: "P k" have 2: "k \ (GREATEST k. P k)" using Greatest_le 1 assms(2) by this show "False" using assms(1) 2 by auto qed lemma finite_set_of_finite_maps': assumes "finite A" "finite B" shows "finite {m. dom m \ A \ ran m \ B}" proof - have "{m. dom m \ A \ ran m \ B} = (\ \ \ Pow A. {m. dom m = \ \ ran m \ B})" by auto also have "finite \" using finite_subset assms by (auto intro: finite_set_of_finite_maps) finally show ?thesis by this qed primrec alternate :: "('a \ 'a) \ ('a \ 'a) \ nat \ ('a \ 'a)" where "alternate f g 0 = id" | "alternate f g (Suc k) = alternate g f k \ f" lemma alternate_Suc[simp]: "alternate f g (Suc k) = (if even k then f else g) \ alternate f g k" proof (induct k arbitrary: f g) case (0) show ?case by simp next case (Suc k) have "alternate f g (Suc (Suc k)) = alternate g f (Suc k) \ f" by auto also have "\ = (if even k then g else f) \ (alternate g f k \ f)" unfolding Suc by auto also have "\ = (if even (Suc k) then f else g) \ alternate f g (Suc k)" by auto finally show ?case by this qed declare alternate.simps(2)[simp del] lemma alternate_antimono: assumes "\ x. f x \ x" "\ x. g x \ x" shows "antimono (alternate f g)" proof fix k l :: nat assume 1: "k \ l" obtain n where 2: "l = k + n" using le_Suc_ex 1 by auto have 3: "alternate f g (k + n) \ alternate f g k" proof (induct n) case (0) show ?case by simp next case (Suc n) have "alternate f g (k + Suc n) \ alternate f g (k + n)" using assms by (auto intro: le_funI) also have "\ \ alternate f g k" using Suc by this finally show ?case by this qed show "alternate f g l \ alternate f g k" using 3 unfolding 2 by this qed end diff --git a/thys/QR_Decomposition/Gram_Schmidt.thy b/thys/QR_Decomposition/Gram_Schmidt.thy --- a/thys/QR_Decomposition/Gram_Schmidt.thy +++ b/thys/QR_Decomposition/Gram_Schmidt.thy @@ -1,913 +1,913 @@ (* Title: Gram_Schmidt.thy Author: Jose Divasón Author: Jesús Aransay *) section\The Gram-Schmidt algorithm\ theory Gram_Schmidt imports Miscellaneous_QR Projections begin subsection\Gram-Schmidt algorithm\ text\ The algorithm is used to orthogonalise a set of vectors. The Gram-Schmidt process takes a set of vectors \S\ and generates another orthogonal set that spans the same subspace as \S\. We present three ways to compute the Gram-Schmidt algorithm. \begin{enumerate} \item The fist one has been developed thinking about the simplicity of its formalisation. Given a list of vectors, the output is another list of orthogonal vectors with the same span. Such a list is constructed following the Gram-Schmidt process presented in any book, but in the reverse order (starting the process from the last element of the input list). \item Based on previous formalization, another function has been defined to compute the process of the Gram-Schmidt algorithm in the natural order (starting from the first element of the input list). \item The third way has as input and output a matrix. The algorithm is applied to the columns of a matrix, obtaining a matrix whose columns are orthogonal and where the column space is kept. This will be a previous step to compute the QR decomposition. \end{enumerate} Every function can be executed with arbitrary precision (using rational numbers). \ subsubsection\First way\ definition Gram_Schmidt_step :: "('a::{real_inner}^'b) => ('a^'b) list => ('a^'b) list" where "Gram_Schmidt_step a ys = ys @ [(a - proj_onto a (set ys))]" definition "Gram_Schmidt xs = foldr Gram_Schmidt_step xs []" lemma Gram_Schmidt_cons: "Gram_Schmidt (a#xs) = Gram_Schmidt_step a (Gram_Schmidt xs)" unfolding Gram_Schmidt_def by auto lemma basis_orthogonal': fixes xs::"('a::{real_inner}^'b) list" shows "length (Gram_Schmidt xs) = length (xs) \ span (set (Gram_Schmidt xs)) = span (set xs) \ pairwise orthogonal (set (Gram_Schmidt xs))" proof (induct xs) case Nil show ?case unfolding Gram_Schmidt_def pairwise_def by fastforce next case (Cons a xs) have l: "length (Gram_Schmidt xs) = length xs" and s: "span (set (Gram_Schmidt xs)) = span (set xs)" and p: "pairwise orthogonal (set (Gram_Schmidt xs))" using Cons.hyps by auto show "length (Gram_Schmidt (a # xs)) = length (a # xs) \ span (set (Gram_Schmidt (a # xs))) = span (set (a # xs)) \ pairwise orthogonal (set (Gram_Schmidt (a # xs)))" proof show "length (Gram_Schmidt (a # xs)) = length (a # xs)" unfolding Gram_Schmidt_cons unfolding Gram_Schmidt_step_def using l by auto show "span (set (Gram_Schmidt (a # xs))) = span (set (a # xs)) \ pairwise orthogonal (set (Gram_Schmidt (a # xs)))" proof have set_rw1: "set (a # xs) = insert a (set xs)" by simp have set_rw2: "set (Gram_Schmidt (a # xs)) = (insert (a - (\x\set (Gram_Schmidt xs). (a \ x / (x \ x)) *\<^sub>R x)) (set (Gram_Schmidt xs)))" unfolding Gram_Schmidt_cons Gram_Schmidt_step_def proj_onto_def proj_def[abs_def] by auto define C where "C = set (Gram_Schmidt xs)" have finite_C: "finite C" unfolding C_def by auto { fix x k have th0: "!!(a::'a^'b) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps) have "x - k *\<^sub>R (a - (\x\C. (a \ x / (x \ x)) *\<^sub>R x)) \ span C \ x - k *\<^sub>R a \ span C" apply (simp only: scaleR_right_diff_distrib th0) apply (rule span_add_eq) apply (rule span_mul) apply (rule span_sum) apply (rule span_mul) apply (rule span_base) apply assumption done } then show "span (set (Gram_Schmidt (a # xs))) = span (set (a # xs))" unfolding set_eq_iff set_rw2 set_rw1 span_breakdown_eq C_def s[symmetric] by auto with p show "pairwise orthogonal (set (Gram_Schmidt (a # xs)))" using pairwise_orthogonal_proj_set[OF finite_C] unfolding set_rw2 unfolding C_def proj_def[symmetric] proj_onto_def[symmetric] by simp qed qed qed lemma card_Gram_Schmidt: fixes xs::"('a::{real_inner}^'b) list" assumes "distinct xs" shows "card(set (Gram_Schmidt xs)) \ card (set (xs))" using assms proof (induct xs) case Nil show ?case unfolding Gram_Schmidt_def by simp next case (Cons x xs) define b where "b = (\xa\set (Gram_Schmidt xs). (x \ xa / (xa \ xa)) *\<^sub>R xa)" have distinct_xs: "distinct xs" using Cons.prems by auto show ?case proof (cases "x - b \ set (Gram_Schmidt xs)") case True have "card (set (Gram_Schmidt (x # xs))) = card (insert (x - b) (set (Gram_Schmidt xs)))" unfolding Gram_Schmidt_cons Gram_Schmidt_step_def b_def unfolding proj_onto_def proj_def[abs_def] by simp also have "... = Suc (card (set (Gram_Schmidt xs)))" proof (rule card_insert_disjoint) show "finite (set (Gram_Schmidt xs))" by simp show "x - b \ set (Gram_Schmidt xs)" using True . qed also have "... \ Suc (card (set xs))" using Cons.hyps[OF distinct_xs] by simp also have "... = Suc (length (remdups xs))" unfolding List.card_set .. also have "... \ (length (remdups (x # xs)))" by (metis Cons.prems distinct_xs impossible_Cons not_less_eq_eq remdups_id_iff_distinct) also have "... \ (card (set (x # xs)))" by (metis dual_order.refl length_remdups_card_conv) finally show ?thesis . next case False have x_b_in: "x - b \ set (Gram_Schmidt xs)" using False by simp have "card (set (Gram_Schmidt (x # xs))) = card (insert (x - b) (set (Gram_Schmidt xs)))" unfolding Gram_Schmidt_cons Gram_Schmidt_step_def b_def unfolding proj_onto_def proj_def[abs_def] by simp also have "... = (card (set (Gram_Schmidt xs)))" by (metis False insert_absorb) also have "... \ (card (set xs))" using Cons.hyps[OF distinct_xs] . also have "... \ (card (set (x # xs)))" unfolding List.card_set by simp finally show ?thesis . qed qed lemma orthogonal_basis_exists: fixes V :: "(real^'b) list" assumes B: "is_basis (set V)" and d: "distinct V" shows "vec.independent (set (Gram_Schmidt V)) \ (set V) \ vec.span (set (Gram_Schmidt V)) \ (card (set (Gram_Schmidt V)) = vec.dim (set V)) \ pairwise orthogonal (set (Gram_Schmidt V))" proof - have "(set V) \ vec.span (set (Gram_Schmidt V))" using basis_orthogonal'[of V] using vec.span_superset[where ?'a=real, where ?'b='b] by (auto simp: span_vec_eq) moreover have "pairwise orthogonal (set (Gram_Schmidt V))" using basis_orthogonal'[of V] by blast moreover have c: "(card (set (Gram_Schmidt V)) = vec.dim (set V))" proof - have card_eq_dim: "card (set V) = vec.dim (set V)" by (metis B independent_is_basis vec.dim_span vec.indep_card_eq_dim_span) have "vec.dim (set V) \ (card (set (Gram_Schmidt V)))" using B unfolding is_basis_def using vec.independent_span_bound[of "(set (Gram_Schmidt V))" "set V"] using basis_orthogonal'[of V] by (simp add: calculation(1) local.card_eq_dim) moreover have "(card (set (Gram_Schmidt V))) \ vec.dim (set V)" using card_Gram_Schmidt[OF d] card_eq_dim by auto ultimately show ?thesis by auto qed moreover have "vec.independent (set (Gram_Schmidt V))" proof (rule vec.card_le_dim_spanning[of _ "UNIV::(real^'b) set"]) show "set (Gram_Schmidt V) \ (UNIV::(real^'b) set)" by simp show "UNIV \ vec.span (set (Gram_Schmidt V))" using basis_orthogonal'[of V] using B unfolding is_basis_def by (simp add: span_vec_eq) show "finite (set (Gram_Schmidt V))" by simp show "card (set (Gram_Schmidt V)) \ vec.dim (UNIV::(real^'b) set)" by (metis c top_greatest vec.dim_subset) qed ultimately show ?thesis by simp qed corollary orthogonal_basis_exists': fixes V :: "(real^'b) list" assumes B: "is_basis (set V)" and d: "distinct V" shows "is_basis (set (Gram_Schmidt V)) \ distinct (Gram_Schmidt V) \ pairwise orthogonal (set (Gram_Schmidt V))" using B orthogonal_basis_exists basis_orthogonal' card_distinct d vec.dim_unique distinct_card is_basis_def subset_refl by (metis span_vec_eq) subsubsection\Second way\ text\This definition applies the Gram Schmidt process starting from the first element of the list.\ definition "Gram_Schmidt2 xs = Gram_Schmidt (rev xs)" lemma basis_orthogonal2: fixes xs::"('a::{real_inner}^'b) list" shows "length (Gram_Schmidt2 xs) = length (xs) \ span (set (Gram_Schmidt2 xs)) = span (set xs) \ pairwise orthogonal (set (Gram_Schmidt2 xs))" by (metis Gram_Schmidt2_def basis_orthogonal' length_rev set_rev) lemma card_Gram_Schmidt2: fixes xs::"('a::{real_inner}^'b) list" assumes "distinct xs" shows "card(set (Gram_Schmidt2 xs)) \ card (set (xs))" by (metis Gram_Schmidt2_def assms card_Gram_Schmidt distinct_rev set_rev) lemma orthogonal_basis_exists2: fixes V :: "(real^'b) list" assumes B: "is_basis (set V)" and d: "distinct V" shows "vec.independent (set (Gram_Schmidt2 V)) \ (set V) \ vec.span (set (Gram_Schmidt2 V)) \ (card (set (Gram_Schmidt2 V)) = vec.dim (set V)) \ pairwise orthogonal (set (Gram_Schmidt2 V))" by (metis Gram_Schmidt.orthogonal_basis_exists Gram_Schmidt2_def distinct_rev set_rev B basis_orthogonal2 d) subsubsection\Third way\ text\The following definitions applies the Gram Schmidt process in the columns of a given matrix. It is previous step to the computation of the QR decomposition.\ definition Gram_Schmidt_column_k :: "'a::{real_inner}^'cols::{mod_type}^'rows \ nat \ 'a^'cols::{mod_type}^'rows" where "Gram_Schmidt_column_k A k = (\ a. (\ b. (if b = from_nat k then (column b A - (proj_onto (column b A) {column i A|i. i < b})) else (column b A)) $ a))" definition "Gram_Schmidt_upt_k A k = foldl Gram_Schmidt_column_k A [0..<(Suc k)]" definition "Gram_Schmidt_matrix A = Gram_Schmidt_upt_k A (ncols A - 1)" text\Some definitions and lemmas in order to get execution.\ definition "Gram_Schmidt_column_k_row A k a = vec_lambda(\b. (if b = from_nat k then (column b A - (\x\{column i A|i. i < b}. ((column b A) \ x / (x \ x)) *\<^sub>R x)) else (column b A)) $ a)" lemma Gram_Schmidt_column_k_row_code[code abstract]: "vec_nth (Gram_Schmidt_column_k_row A k a) = (%b. (if b = from_nat k then (column b A - (\x\{column i A|i. i < b}. ((column b A) \ x / (x \ x)) *\<^sub>R x)) else (column b A)) $ a)" unfolding Gram_Schmidt_column_k_row_def by (metis (lifting) vec_lambda_beta) lemma Gram_Schmidt_column_k_code[code abstract]: "vec_nth (Gram_Schmidt_column_k A k) = Gram_Schmidt_column_k_row A k" unfolding Gram_Schmidt_column_k_def unfolding Gram_Schmidt_column_k_row_def[abs_def] unfolding proj_onto_def proj_def[abs_def] by fastforce text\Proofs\ lemma Gram_Schmidt_upt_k_suc: "Gram_Schmidt_upt_k A (Suc k) = (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k))" proof - have rw: "[0..(Suc k)} = {column i (Gram_Schmidt_upt_k A k) |i. to_nat i\k} \ {column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) - (\x\{column i (Gram_Schmidt_upt_k A k) |i. to_nat i\k}. (x \ (column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k)) / (x \ x)) *\<^sub>R x)}" proof - have set_rw: "{\ ia. Gram_Schmidt_upt_k A k $ ia $ i |i. i < from_nat (Suc k)} = {\ ia. Gram_Schmidt_upt_k A k $ ia $ i |i. to_nat i \ k}" proof (auto, vector, metis less_Suc_eq_le to_nat_le) fix i::'cols assume "to_nat i \ k" hence "to_nat i < Suc k" by simp hence i_less_suc: "i < from_nat (Suc k)" using from_nat_le[OF _ k] by simp show "\l. (\j. Gram_Schmidt_upt_k A k $ j $ i) = (\j'. Gram_Schmidt_upt_k A k $ j' $ l) \ l < mod_type_class.from_nat (Suc k)" by (rule exI[of _ i], simp add: i_less_suc) qed have rw: "[0..(Suc k)} = {column i (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k)) |i. to_nat i \ Suc k}" unfolding Gram_Schmidt_upt_k_def unfolding rw unfolding foldl_append unfolding foldl.simps .. also have "... = {column i (Gram_Schmidt_upt_k A k) |i. to_nat i\k} \ {column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) - (\x\{column i (Gram_Schmidt_upt_k A k) |i. to_nat i\k}. (x \ (column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k)) / (x \ x)) *\<^sub>R x)}" proof (auto) fix i::'cols assume ik: "to_nat i \ k" show "\ia. column i (Gram_Schmidt_upt_k A k) = column ia (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k)) \ to_nat ia \ Suc k" proof (rule exI[of _ i], rule conjI) have i_less_suc: "to_nat i < Suc k" using ik by auto thus "to_nat i \ Suc k" by simp show "column i (Gram_Schmidt_upt_k A k) = column i (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k))" using column_Gram_Schmidt_upt_k_preserves[OF i_less_suc k, of A] unfolding Gram_Schmidt_upt_k_suc .. qed next show "\a. column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) - (\x\{column i (Gram_Schmidt_upt_k A k) |i. to_nat i \ k}. (x \ column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) / (x \ x)) *\<^sub>R x) = column a (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k)) \ to_nat a \ Suc k" proof (rule exI[of _ "from_nat (Suc k)::'cols"], rule conjI) show "to_nat (from_nat (Suc k)::'cols) \ Suc k" unfolding to_nat_from_nat_id[OF k] .. show "column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) - (\x\{column i (Gram_Schmidt_upt_k A k) |i. to_nat i \ k}. (x \ column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) / (x \ x)) *\<^sub>R x) = column (from_nat (Suc k)) (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k))" unfolding Gram_Schmidt_column_k_def column_def apply auto unfolding set_rw unfolding vector_scaleR_component[symmetric] unfolding sum_component[symmetric] unfolding proj_onto_def proj_def[abs_def] unfolding proj_onto_sum_rw by vector qed next fix i assume col_not_eq: "column i (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k)) \ column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) - (\x\{column i (Gram_Schmidt_upt_k A k) |i. to_nat i \ k}. (x \ column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) / (x \ x)) *\<^sub>R x)" and i: "to_nat i \ Suc k" have i_not_suc_k: "i \ from_nat (Suc k)" proof (rule ccontr, simp) assume i2: "i = from_nat (Suc k)" have "column i (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k)) = column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) - (\x\{column i (Gram_Schmidt_upt_k A k) |i. to_nat i \ k}. (x \ column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) / (x \ x)) *\<^sub>R x)" unfolding i2 Gram_Schmidt_column_k_def column_def apply auto unfolding set_rw unfolding vector_scaleR_component[symmetric] unfolding sum_component[symmetric] unfolding proj_onto_def proj_def[abs_def] unfolding proj_onto_sum_rw by vector thus False using col_not_eq by contradiction qed show "\ia. column i (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k)) = column ia (Gram_Schmidt_upt_k A k) \ to_nat ia \ k" proof (rule exI[of _ i], rule conjI, unfold Gram_Schmidt_upt_k_suc[symmetric], rule column_Gram_Schmidt_upt_k_preserves) show "to_nat i < Suc k" using i i_not_suc_k by (metis le_imp_less_or_eq from_nat_to_nat_id) thus "to_nat i \ k" using less_Suc_eq_le by simp show "Suc k < CARD('cols)" using k . qed qed finally show ?thesis . qed lemma orthogonal_Gram_Schmidt_upt_k: assumes s: "k < ncols A" shows "pairwise orthogonal ({column i (Gram_Schmidt_upt_k A k) |i. to_nat i\k})" using s proof (induct k) case 0 have set_rw: "{column i (Gram_Schmidt_upt_k A 0) |i. to_nat i \ 0} = {column 0 (Gram_Schmidt_upt_k A 0)}" by (simp add: to_nat_eq_0) show ?case unfolding set_rw unfolding pairwise_def by auto next case (Suc k) have rw: "[0.. UNIV} = {column i (Gram_Schmidt_upt_k A (ncols A - 1)) |i. to_nat i\ (ncols A - 1)}" proof (auto) fix i show "\ia. column i (Gram_Schmidt_matrix A) = column ia (Gram_Schmidt_upt_k A (ncols A - Suc 0)) \ to_nat ia \ ncols A - Suc 0" apply (rule exI[of _ i]) unfolding Gram_Schmidt_matrix_def using to_nat_less_card[of i] unfolding ncols_def by fastforce show "\ia. column i (Gram_Schmidt_upt_k A (ncols A - Suc 0)) = column ia (Gram_Schmidt_matrix A)" unfolding Gram_Schmidt_matrix_def by auto qed corollary orthogonal_Gram_Schmidt_matrix: shows "pairwise orthogonal ({column i (Gram_Schmidt_matrix A) |i. i \ UNIV})" unfolding columns_Gram_Schmidt_matrix_rw by (rule orthogonal_Gram_Schmidt_upt_k, simp add: ncols_def) corollary orthogonal_Gram_Schmidt_matrix2: shows "pairwise orthogonal (columns (Gram_Schmidt_matrix A))" unfolding columns_def using orthogonal_Gram_Schmidt_matrix . lemma column_Gram_Schmidt_column_k: fixes A::"'a::{real_inner}^'n::{mod_type}^'m::{mod_type}" shows "column k (Gram_Schmidt_column_k A (to_nat k)) = (column k A) - (\x\{column i A|i. i < k}. (x \ (column k A) / (x \ x)) *\<^sub>R x)" unfolding Gram_Schmidt_column_k_def column_def unfolding from_nat_to_nat_id unfolding proj_onto_def proj_def[abs_def] unfolding proj_onto_sum_rw by vector lemma column_Gram_Schmidt_column_k': fixes A::"'a::{real_inner}^'n::{mod_type}^'m::{mod_type}" assumes i_not_k: "i\k" shows "column i (Gram_Schmidt_column_k A (to_nat k)) = (column i A)" using i_not_k unfolding Gram_Schmidt_column_k_def column_def unfolding from_nat_to_nat_id by vector definition "cols_upt_k A k = {column i A|i. i\from_nat k}" lemma cols_upt_k_insert: fixes A::"'a^'n::{mod_type}^'m::{mod_type}" assumes k: "(Suc k) from_nat (Suc k)" and "column i A \ column (from_nat (Suc k)) A" hence i_not_suc_k: "i \ from_nat (Suc k)" by auto have i_le: "i \ from_nat k" proof - have "i < from_nat (Suc k)" by (metis le_imp_less_or_eq i i_not_suc_k) thus ?thesis by (metis Suc_eq_plus1 from_nat_suc le_Suc not_less) qed thus "\ia. column i A = column ia A \ ia \ from_nat k" by auto next fix i::'n assume i: "i \ from_nat k" also have "... < from_nat (Suc k)" by (rule from_nat_mono[OF _ k[unfolded ncols_def]], simp) finally have "i \ from_nat (Suc k)" by simp thus "\ia. column i A = column ia A \ ia \ from_nat (Suc k)" by auto qed lemma columns_eq_cols_upt_k: fixes A::"'a^'cols::{mod_type}^'rows::{mod_type}" shows "cols_upt_k A (ncols A - 1) = columns A" proof (unfold cols_upt_k_def columns_def, auto) fix i show "\ia. column i A = column ia A \ ia \ from_nat (ncols A - Suc 0)" proof (rule exI[of _ i], simp) have "to_nat i < ncols A" using to_nat_less_card[of i] unfolding ncols_def by simp hence "to_nat i \ (ncols A - 1)" by simp hence "to_nat i \ to_nat (from_nat (ncols A - 1)::'cols)" using to_nat_from_nat_id[of "ncols A - 1", where ?'a='cols] unfolding ncols_def by simp thus "i \ from_nat (ncols A - Suc 0)" by (metis One_nat_def less_le_not_le linear to_nat_mono) qed qed lemma span_cols_upt_k_Gram_Schmidt_column_k: fixes A::"'a::{real_inner}^'n::{mod_type}^'m::{mod_type}" assumes "k < ncols A" and "j < ncols A" shows "span (cols_upt_k A k) = span (cols_upt_k (Gram_Schmidt_column_k A j) k)" using assms proof (induct k) case 0 have set_rw: "{\ ia. A $ ia $ i |i. i < 0} = {}" using least_mod_type not_less by auto have set_rw2: "{column i (Gram_Schmidt_column_k A j) |i. i \ 0} = {column 0 (Gram_Schmidt_column_k A j)}" by (auto, metis eq_iff least_mod_type) have set_rw3: "{column i A |i. i \ 0} ={column 0 A}" by (auto, metis eq_iff least_mod_type) have col_0_eq: "column 0 (Gram_Schmidt_column_k A j) = column 0 A" unfolding Gram_Schmidt_column_k_def column_def unfolding proj_onto_def proj_def[abs_def] by (simp add: set_rw) show ?case unfolding cols_upt_k_def from_nat_0 unfolding set_rw2 set_rw3 unfolding col_0_eq .. next case (Suc k) have hyp: "span (cols_upt_k A k) = span (cols_upt_k (Gram_Schmidt_column_k A j) k)" using Suc.prems Suc.hyps by auto have set_rw1: "(cols_upt_k A (Suc k)) = insert (column (from_nat (Suc k)) A) (cols_upt_k A k)" using cols_upt_k_insert by (auto intro!: cols_upt_k_insert[OF Suc.prems(1)]) have set_rw2: "(cols_upt_k (Gram_Schmidt_column_k A j) (Suc k)) = insert (column (from_nat (Suc k)) (Gram_Schmidt_column_k A j)) (cols_upt_k (Gram_Schmidt_column_k A j) k)" using cols_upt_k_insert Suc.prems(1) unfolding ncols_def by auto show ?case proof (cases "j=Suc k") case False have suc_not_k: "from_nat (Suc k) \ (from_nat j::'n)" proof (rule ccontr, simp) assume "from_nat (Suc k) = (from_nat j::'n)" hence "Suc k = j" apply (rule from_nat_eq_imp_eq) using Suc.prems unfolding ncols_def . thus False using False by simp qed have tnfnj: "to_nat (from_nat j::'n) = j" using to_nat_from_nat_id[OF Suc.prems(2)[unfolded ncols_def]] . let ?a_suc_k = "column (from_nat (Suc k)) A" have col_eq: "column (from_nat (Suc k)) (Gram_Schmidt_column_k A j) = ?a_suc_k" using column_Gram_Schmidt_column_k'[OF suc_not_k] unfolding tnfnj . have k: "k from_nat k} = {column i A |i. i < from_nat (Suc k)}" proof (auto) fix i::'n assume i: "i \ from_nat k" also have "... < from_nat (Suc k)" by (rule from_nat_mono[OF _ Suc.prems(1)[unfolded ncols_def]], simp) finally show "\ia. column i A = column ia A \ ia < from_nat (Suc k)" by auto next fix i::'n assume i: "i < from_nat (Suc k)" hence "i\ from_nat k" unfolding Suc_eq_plus1 unfolding from_nat_suc by (metis le_Suc not_less) thus " \ia. column i A = column ia A \ ia \ from_nat k" by auto qed have rw: "column (from_nat (Suc k)) (Gram_Schmidt_column_k A j) = (a - (\x\cols_upt_k A k. (x \ a / (x \ x)) *\<^sub>R x))" unfolding Gram_Schmidt_column_k_def True unfolding cols_upt_k_def a_def C_def unfolding column_def apply auto unfolding column_def[symmetric] col_rw minus_vec_def unfolding column_def vec_lambda_beta unfolding proj_onto_def proj_def[abs_def] unfolding proj_onto_sum_rw by auto have finite_C: "finite C" unfolding C_def cols_upt_k_def by auto { fix x b have th0: "!!(a::'a^'m::{mod_type}) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps) have "x - b *\<^sub>R (a - (\x\C. (x \ a / (x \ x)) *\<^sub>R x)) \ span C \ x - b *\<^sub>R a \ span C" apply (simp only: scaleR_right_diff_distrib th0) apply (rule span_add_eq) apply (rule span_mul) apply (rule span_sum) apply (rule span_mul) apply (rule span_base) apply assumption done } thus ?thesis unfolding set_eq_iff unfolding C_def B_def unfolding set_rw1 unfolding set_rw2 unfolding span_breakdown_eq unfolding hyp by (metis (mono_tags) B_def a_def rw) qed qed corollary span_Gram_Schmidt_column_k: fixes A::"'a::{real_inner}^'n::{mod_type}^'m::{mod_type}" assumes "k ia. A $ ia $ i |i. i < from_nat 0} = {}" by (auto, metis less_le_not_le least_mod_type from_nat_0) thus "\ia. column i (Gram_Schmidt_column_k A 0) = column ia A" unfolding Gram_Schmidt_column_k_def column_def unfolding proj_onto_def proj_def[abs_def] by auto show "\ia. column i A = column ia (Gram_Schmidt_column_k A 0)" using set_rw unfolding Gram_Schmidt_column_k_def column_def unfolding from_nat_0 unfolding proj_onto_def proj_def[abs_def] by force qed thus ?case unfolding Gram_Schmidt_upt_k_def by auto next case (Suc k) have hyp: "span (columns A) = span (columns (Gram_Schmidt_upt_k A k))" using Suc.prems Suc.hyps by auto have "span (columns (Gram_Schmidt_upt_k A (Suc k))) = span (columns (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k)))" unfolding Gram_Schmidt_upt_k_suc .. also have "... = span (columns (Gram_Schmidt_upt_k A k))" using span_Gram_Schmidt_column_k[of "Suc k" "(Gram_Schmidt_upt_k A k)"] using Suc.prems unfolding ncols_def by auto finally show ?case using hyp by auto qed corollary span_Gram_Schmidt_matrix: fixes A::"'a::{real_inner}^'n::{mod_type}^'m::{mod_type}" shows "span (columns A) = span (columns (Gram_Schmidt_matrix A))" unfolding Gram_Schmidt_matrix_def by (simp add: span_Gram_Schmidt_upt_k ncols_def) lemma is_basis_columns_Gram_Schmidt_matrix: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes b: "is_basis (columns A)" and c: "card (columns A) = ncols A" shows "is_basis (columns (Gram_Schmidt_matrix A)) \ card (columns (Gram_Schmidt_matrix A)) = ncols A" proof - have span_UNIV: "vec.span (columns (Gram_Schmidt_matrix A)) = (UNIV::(real^'m::{mod_type}) set)" using span_Gram_Schmidt_matrix b unfolding is_basis_def by (metis span_vec_eq) moreover have c_eq: "card (columns (Gram_Schmidt_matrix A)) = ncols A" proof - have "card (columns A) \ card (columns (Gram_Schmidt_matrix A))" by (metis b is_basis_def finite_columns vec.independent_span_bound span_UNIV top_greatest) thus ?thesis using c using card_columns_le_ncols by (metis le_antisym ncols_def) qed moreover have "vec.independent (columns (Gram_Schmidt_matrix A))" proof (rule vec.card_le_dim_spanning[of _ "UNIV::(real^'m::{mod_type}) set"]) show "columns (Gram_Schmidt_matrix A) \ UNIV" by simp show "UNIV \ vec.span (columns (Gram_Schmidt_matrix A))" using span_UNIV by auto show "finite (columns (Gram_Schmidt_matrix A))" using finite_columns . show "card (columns (Gram_Schmidt_matrix A)) \ vec.dim (UNIV::(real^'m::{mod_type}) set)" by (metis b c c_eq eq_iff is_basis_def vec.dim_span_eq_card_independent) qed ultimately show ?thesis unfolding is_basis_def by simp qed text\From here on, we present some lemmas that will be useful for the formalisation of the QR decomposition.\ lemma column_gr_k_Gram_Schmidt_upt: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes "i>k" and i: "i 0" unfolding from_nat_0[symmetric] using bij_from_nat[where ?'a='n] unfolding bij_betw_def by (metis from_nat_eq_imp_eq gr_implies_not0 i ncols_def neq0_conv) thus "column (from_nat i) (Gram_Schmidt_upt_k A 0) = column (from_nat i) A" unfolding Gram_Schmidt_upt_k_def by (simp add: Gram_Schmidt_column_k_def from_nat_0 column_def) next case (Suc k) have hyp: "column (from_nat i) (Gram_Schmidt_upt_k A k) = column (from_nat i) A" using Suc.hyps Suc.prems by auto have to_nat_from_nat_suc_k: "(to_nat (from_nat (Suc k)::'n)) = Suc k" by (metis Suc.prems dual_order.strict_trans from_nat_not_eq i ncols_def) have "column (from_nat i) (Gram_Schmidt_upt_k A (Suc k)) = column (from_nat i) (Gram_Schmidt_column_k (Gram_Schmidt_upt_k A k) (Suc k))" unfolding Gram_Schmidt_upt_k_suc .. also have "... = column (from_nat i) (Gram_Schmidt_upt_k A k)" proof (rule column_Gram_Schmidt_column_k' [of "from_nat i" "from_nat (Suc k)" " (Gram_Schmidt_upt_k A k)", unfolded to_nat_from_nat_suc_k]) show "from_nat i \ (from_nat (Suc k)::'n)" by (metis Suc.prems not_less_iff_gr_or_eq from_nat_not_eq i ncols_def to_nat_from_nat_suc_k) qed finally show ?case unfolding hyp . qed lemma columns_Gram_Schmidt_upt_k_rw: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes k: "Suc k < ncols A" shows "{column i (Gram_Schmidt_upt_k A (Suc k)) |i. i < from_nat (Suc k)} = {column i (Gram_Schmidt_upt_k A k) |i. i < from_nat (Suc k)}" proof (auto) fix i::'n assume i: "i < from_nat (Suc k)" have to_nat_from_nat_k: "to_nat (from_nat (Suc k)::'n) = Suc k" using to_nat_from_nat_id k unfolding ncols_def by auto show "\ia. column i (Gram_Schmidt_upt_k A (Suc k)) = column ia (Gram_Schmidt_upt_k A k) \ ia < from_nat (Suc k)" by (metis column_Gram_Schmidt_upt_k_preserves i k ncols_def to_nat_le) show "\ia. column i (Gram_Schmidt_upt_k A k) = column ia (Gram_Schmidt_upt_k A (Suc k)) \ ia < from_nat (Suc k)" by (metis column_Gram_Schmidt_upt_k_preserves i k ncols_def to_nat_le) qed lemma column_Gram_Schmidt_upt_k: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes "kx\{column i (Gram_Schmidt_upt_k A k)|i. i < (from_nat k)}. (x \ (column (from_nat k) A) / (x \ x)) *\<^sub>R x)" using assms proof (induct k, unfold from_nat_0) have set_rw: "{column i (Gram_Schmidt_upt_k A 0) |i. i < 0} = {}" by (auto, metis least_mod_type not_le) have set_rw2: "{column i A |i. i < 0} = {}" by (auto, metis least_mod_type not_le) have col_rw: "column 0 A = column 0 (Gram_Schmidt_upt_k A 0)" unfolding Gram_Schmidt_upt_k_def apply auto unfolding Gram_Schmidt_column_k_def from_nat_0 unfolding column_def using set_rw2 unfolding proj_onto_def proj_def[abs_def] by vector show "column 0 (Gram_Schmidt_upt_k A 0) = column 0 A - (\x\{column i (Gram_Schmidt_upt_k A 0) |i. i < 0}. (x \ column 0 A / (x \ x)) *\<^sub>R x)" unfolding set_rw col_rw by simp next case (Suc k) let ?ak="column (from_nat k) A" let ?uk="column (from_nat k) (Gram_Schmidt_upt_k A k)" let ?a_suck= "column (from_nat (Suc k)) A" let ?u_suck="column (from_nat (Suc k)) (Gram_Schmidt_upt_k A (Suc k))" have to_nat_from_nat_k: "to_nat (from_nat (Suc k)::'n) = (Suc k)" using to_nat_from_nat_id Suc.prems unfolding ncols_def by auto have a_suc_rw: "column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) = ?a_suck" by (rule column_gr_k_Gram_Schmidt_upt, auto simp add: Suc.prems) have set_rw: "{column i (Gram_Schmidt_upt_k A (Suc k)) |i. i < from_nat (Suc k)} = {column i (Gram_Schmidt_upt_k A k) |i. i < from_nat (Suc k)}" by (rule columns_Gram_Schmidt_upt_k_rw[OF Suc.prems]) have "?u_suck = column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) - (\x\{column i (Gram_Schmidt_upt_k A k) |i. i < from_nat (Suc k)}. (x \ column (from_nat (Suc k)) (Gram_Schmidt_upt_k A k) / (x \ x)) *\<^sub>R x)" unfolding Gram_Schmidt_upt_k_suc using column_Gram_Schmidt_column_k[of "from_nat (Suc k)" "(Gram_Schmidt_upt_k A k)"] unfolding to_nat_from_nat_k by auto also have "... = ?a_suck - (\x\{column i (Gram_Schmidt_upt_k A k) |i. i < from_nat (Suc k)}. (x \ ?a_suck / (x \ x)) *\<^sub>R x)" unfolding a_suc_rw .. finally show "?u_suck = ?a_suck - (\x\{column i (Gram_Schmidt_upt_k A (Suc k)) |i. i < from_nat (Suc k)}. (x \ ?a_suck / (x \ x)) *\<^sub>R x)" unfolding set_rw . qed lemma column_Gram_Schmidt_upt_k_preserves2: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes "a\(from_nat i)" and "i \ j" and "j from_nat (Suc j)" by (rule from_nat_mono'[OF Suc.prems(2) Suc.prems(3)[unfolded ncols_def]]) hence "from_nat i = (from_nat (Suc j)::'n)" using Suc.prems(1) unfolding True by simp hence i_eq_suc: "i=Suc j" apply (rule from_nat_eq_imp_eq) using Suc.prems unfolding ncols_def by auto show ?thesis unfolding True i_eq_suc .. qed qed lemma set_columns_Gram_Schmidt_matrix: fixes A::"real^'n::{mod_type}^'m::{mod_type}" shows "{column i (Gram_Schmidt_matrix A)|i. i < k} = {column i (Gram_Schmidt_upt_k A (to_nat k))|i. i < k}" proof (auto) fix i assume i: "iia. column i (Gram_Schmidt_matrix A) = column ia (Gram_Schmidt_upt_k A (to_nat k)) \ ia < k" proof (rule exI[of _ i], rule conjI) show "column i (Gram_Schmidt_matrix A) = column i (Gram_Schmidt_upt_k A (to_nat k))" unfolding Gram_Schmidt_matrix_def proof (rule column_Gram_Schmidt_upt_k_preserves2[symmetric]) show "i \ from_nat (to_nat k)" using i unfolding from_nat_to_nat_id by auto show "to_nat k \ ncols A - 1" unfolding ncols_def using to_nat_less_card[of k] by auto show "ncols A - 1 < ncols A" unfolding ncols_def by simp qed show "i < k" using i . qed show "\ia. column i (Gram_Schmidt_upt_k A (to_nat k)) = column ia (Gram_Schmidt_matrix A) \ ia < k" proof (rule exI[of _ i], rule conjI) show "column i (Gram_Schmidt_upt_k A (to_nat k)) = column i (Gram_Schmidt_matrix A)" unfolding Gram_Schmidt_matrix_def proof (rule column_Gram_Schmidt_upt_k_preserves2) show "i \ from_nat (to_nat k)" using i unfolding from_nat_to_nat_id by auto show "to_nat k \ ncols A - 1" unfolding ncols_def using to_nat_less_card[of k] by auto show "ncols A - 1 < ncols A" unfolding ncols_def by simp qed show "i < k" using i . qed qed lemma column_Gram_Schmidt_matrix: fixes A::"real^'n::{mod_type}^'m::{mod_type}" shows "column k (Gram_Schmidt_matrix A) = (column k A) - (\x\{column i (Gram_Schmidt_matrix A)|i. i < k}. (x \ (column k A) / (x \ x)) *\<^sub>R x)" proof - have k: "to_nat k < ncols A" using to_nat_less_card[of k] unfolding ncols_def by simp have "column k (Gram_Schmidt_matrix A) = column k (Gram_Schmidt_upt_k A (ncols A - 1))" unfolding Gram_Schmidt_matrix_def .. also have "... = column k (Gram_Schmidt_upt_k A (to_nat k))" proof (rule column_Gram_Schmidt_upt_k_preserves2[symmetric]) show "k \ from_nat (to_nat k)" unfolding from_nat_to_nat_id .. show "to_nat k \ ncols A - 1" unfolding ncols_def using to_nat_less_card[where ?'a='n] - by (metis Nat.le_diff_conv2 add_leE less_diff_conv less_imp_le_nat less_le_not_le + by (metis le_diff_conv2 add_leE less_diff_conv less_imp_le_nat less_le_not_le nat_le_linear suc_not_zero to_nat_plus_one_less_card') show "ncols A - 1 < ncols A" unfolding ncols_def by auto qed also have "... = column k A - (\x\{column i (Gram_Schmidt_upt_k A (to_nat k)) |i. i < k}. (x \ column k A / (x \ x)) *\<^sub>R x)" using column_Gram_Schmidt_upt_k[OF k] unfolding from_nat_to_nat_id by auto also have "... = column k A - (\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. (x \ column k A / (x \ x)) *\<^sub>R x)" unfolding set_columns_Gram_Schmidt_matrix[symmetric] .. finally show ?thesis . qed corollary column_Gram_Schmidt_matrix2: fixes A::"real^'n::{mod_type}^'m::{mod_type}" shows "(column k A) = column k (Gram_Schmidt_matrix A) + (\x\{column i (Gram_Schmidt_matrix A)|i. i < k}. (x \ (column k A) / (x \ x)) *\<^sub>R x)" using column_Gram_Schmidt_matrix[of k A] by simp lemma independent_columns_Gram_Schmidt_matrix: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes b: "vec.independent (columns A)" and c: "card (columns A) = ncols A" shows "vec.independent (columns (Gram_Schmidt_matrix A)) \ card (columns (Gram_Schmidt_matrix A)) = ncols A" using b c card_columns_le_ncols vec.card_eq_dim_span_indep vec.dim_span eq_iff finite_columns vec.independent_span_bound ncols_def span_Gram_Schmidt_matrix by (metis (no_types, lifting) vec.card_ge_dim_independent vec.dim_span_eq_card_independent span_vec_eq) lemma column_eq_Gram_Schmidt_matrix: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" and c: "column i (Gram_Schmidt_matrix A) = column ia (Gram_Schmidt_matrix A)" shows "i = ia" proof (rule ccontr) assume i_not_ia: "i \ ia" have "columns (Gram_Schmidt_matrix A) = (\x. column x (Gram_Schmidt_matrix A))` (UNIV::('n::{mod_type}) set)" unfolding columns_def by auto also have "... = (\x. column x (Gram_Schmidt_matrix A))` ((UNIV::('n::{mod_type}) set)-{ia})" proof (unfold image_def, auto) fix xa show "\x\UNIV - {ia}. column xa (Gram_Schmidt_matrix A) = column x (Gram_Schmidt_matrix A)" proof (cases "xa = ia") case True thus ?thesis using c i_not_ia by (metis DiffI UNIV_I empty_iff insert_iff) next case False thus ?thesis by auto qed qed finally have columns_rw: "columns (Gram_Schmidt_matrix A) = (\x. column x (Gram_Schmidt_matrix A)) ` (UNIV - {ia})" . have "ncols A = card (columns (Gram_Schmidt_matrix A))" by (metis full_rank_imp_is_basis2 independent_columns_Gram_Schmidt_matrix r) also have "... \ card (UNIV - {ia})" unfolding columns_rw by (rule card_image_le, simp) also have "... = card (UNIV::'n set) - 1" by (simp add: card_Diff_singleton) finally show False unfolding ncols_def - by (metis Nat.add_0_right Nat.le_diff_conv2 One_nat_def Suc_n_not_le_n add_Suc_right one_le_card_finite) + by (metis Nat.add_0_right le_diff_conv2 One_nat_def Suc_n_not_le_n add_Suc_right one_le_card_finite) qed lemma scaleR_columns_Gram_Schmidt_matrix: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes "i \ j" and "rank A = ncols A" shows "column j (Gram_Schmidt_matrix A) \ column i (Gram_Schmidt_matrix A) = 0" proof - have "column j (Gram_Schmidt_matrix A) \ column i (Gram_Schmidt_matrix A)" using column_eq_Gram_Schmidt_matrix assms by auto thus ?thesis using orthogonal_Gram_Schmidt_matrix2 unfolding pairwise_def orthogonal_def columns_def by blast qed subsubsection\Examples of execution\ text\Code lemma\ lemmas Gram_Schmidt_step_def[unfolded proj_onto_def proj_def[abs_def],code] value "let a = map (list_to_vec::real list=> real^4) [[4,-2,-1,2], [-6,3,4,-8], [5,-5,-3,-4]] in map vec_to_list (Gram_Schmidt a)" value "let a = map (list_to_vec::real list=> real^4) [[4,-2,-1,2], [-6,3,4,-8], [5,-5,-3,-4]] in map vec_to_list (Gram_Schmidt2 a)" value "let A = list_of_list_to_matrix [[4,-2,-1,2], [-6,3,4,-8], [5,-5,-3,-4]]::real^4^3 in matrix_to_list_of_list (Gram_Schmidt_matrix A)" end diff --git a/thys/QR_Decomposition/QR_Decomposition.thy b/thys/QR_Decomposition/QR_Decomposition.thy --- a/thys/QR_Decomposition/QR_Decomposition.thy +++ b/thys/QR_Decomposition/QR_Decomposition.thy @@ -1,754 +1,754 @@ (* Title: QR_Decomposition.thy Author: Jose Divasón Author: Jesús Aransay *) section\QR Decomposition\ theory QR_Decomposition imports Gram_Schmidt begin subsection\The QR Decomposition of a matrix\ text\ First of all, it's worth noting what an orthogonal matrix is. In linear algebra, an orthogonal matrix is a square matrix with real entries whose columns and rows are orthogonal unit vectors. Although in some texts the QR decomposition is presented over square matrices, it can be applied to any matrix. There are some variants of the algorithm, depending on the properties that the output matrices satisfy (see for instance, @{url "http://inst.eecs.berkeley.edu/~ee127a/book/login/l_mats_qr.html"}). We present two of them below. Let A be a matrix with m rows and n columns (A is \m \ n\). Case 1: Starting with a matrix whose column rank is maximum. We can define the QR decomposition to obtain: \begin{itemize} \item @{term "A = Q ** R"}. \item Q has m rows and n columns. Its columns are orthogonal unit vectors and @{term "(transpose Q) * Q = mat 1"}. In addition, if A is a square matrix, then Q will be an orthonormal matrix. \item R is \n \ n\, invertible and upper triangular. \end{itemize} Case 2: The called full QR decomposition. We can obtain: \begin{itemize} \item @{term "A = Q ** R"} \item Q is an orthogonal matrix (Q is \m \ m\). \item R is \m \ n\ and upper triangular, but it isn't invertible. \end{itemize} We have decided to formalise the first one, because it's the only useful for solving the linear least squares problem (@{url "http://math.mit.edu/linearalgebra/ila0403.pdf"}). If we have an unsolvable system \A *v x = b\, we can try to find an approximate solution. A plausible choice (not the only one) is to seek an \x\ with the property that \\A ** x - y\\ (the magnitude of the error) is as small as possible. That \x\ is the least squares approximation. We will demostrate that the best approximation (the solution for the linear least squares problem) is the \x\ that satisfies: \(transpose A) ** A *v x = (transpose A) *v b\ Now we want to compute that x. If we are working with the first case, \A\ can be substituted by \Q**R\ and then obtain the solution of the least squares approximation by means of the QR decomposition: \x = (inverse R)**(transpose Q) *v b\ On the contrary, if we are working with the second case after substituting \A\ by \Q**R\ we obtain: \(transpose R) ** R *v x = (transpose R) ** (transpose Q) *v b\ But the \R\ matrix is not invertible (so neither is \transpose R\). The left part of the equation \(transpose R) ** R\ is not going to be an upper triangular matrix, so it can't either be solved using backward-substitution. \ subsubsection\Divide a vector by its norm\ text\An orthogonal matrix is a matrix whose rows (and columns) are orthonormal vectors. So, in order to obtain the QR decomposition, we have to normalise (divide by the norm) the vectors obtained with the Gram-Schmidt algorithm.\ definition "divide_by_norm A = (\ a b. normalize (column b A) $ a)" text\Properties\ lemma norm_column_divide_by_norm: fixes A::"'a::{real_inner}^'cols^'rows" assumes a: "column a A \ 0" shows "norm (column a (divide_by_norm A)) = 1" proof - have not_0: "norm (\ i. A $ i $ a) \ 0" by (metis a column_def norm_eq_zero) have "column a (divide_by_norm A) = (\ i. (1 / norm (\ i. A $ i $ a)) *\<^sub>R A $ i $ a)" unfolding divide_by_norm_def column_def normalize_def by auto also have "... = (1 / norm (\ i. A $ i $ a)) *\<^sub>R (\ i. A $ i $ a)" unfolding vec_eq_iff by auto finally have "norm (column a (divide_by_norm A)) = norm ((1 / norm (\ i. A $ i $ a)) *\<^sub>R (\ i. A $ i $ a))" by simp also have "... = \1 / norm (\ i. A $ i $ a)\ * norm (\ i. A $ i $ a)" unfolding norm_scaleR .. also have "... = (1 / norm (\ i. A $ i $ a)) * norm (\ i. A $ i $ a)" by auto also have "... = 1" using not_0 by auto finally show ?thesis . qed lemma span_columns_divide_by_norm: shows "span (columns A) = span (columns (divide_by_norm A))" unfolding real_vector.span_eq proof (auto) fix x assume x: "x \ columns (divide_by_norm A)" from this obtain i where x_col_i: "x=column i (divide_by_norm A)" unfolding columns_def by blast also have "... = (1/norm (column i A)) *\<^sub>R (column i A)" unfolding divide_by_norm_def column_def normalize_def by vector finally have x_eq: "x=(1/norm (column i A)) *\<^sub>R (column i A)" . show "x \ span (columns A)" by (unfold x_eq, rule span_mul, rule span_base, auto simp add: columns_def) next fix x assume x: "x \ columns A" show "x \ span (columns (divide_by_norm A))" proof (cases "x=0") case True show ?thesis by (metis True span_0) next case False from x obtain i where x_col_i: "x=column i A" unfolding columns_def by blast have "x=column i A" using x_col_i . also have "... = norm (column i A) *\<^sub>R column i (divide_by_norm A)" using False unfolding x_col_i columns_def divide_by_norm_def column_def normalize_def by vector finally have x_eq: "x = norm (column i A) *\<^sub>R column i (divide_by_norm A)" . show "x \ span (columns (divide_by_norm A))" by (unfold x_eq, rule span_mul, rule span_base, auto simp add: columns_def Let_def) qed qed text\Code lemmas\ definition "divide_by_norm_row A a = vec_lambda(% b. ((1 / norm (column b A)) *\<^sub>R column b A) $ a)" lemma divide_by_norm_row_code[code abstract]: "vec_nth (divide_by_norm_row A a) = (% b. ((1 / norm (column b A)) *\<^sub>R column b A) $ a)" unfolding divide_by_norm_row_def by (metis (lifting) vec_lambda_beta) lemma divide_by_norm_code [code abstract]: "vec_nth (divide_by_norm A) = divide_by_norm_row A" unfolding divide_by_norm_def unfolding divide_by_norm_row_def[abs_def] unfolding normalize_def by fastforce subsubsection\The QR Decomposition\ text\The QR decomposition. Given a real matrix @{term "A"}, the algorithm will return a pair @{term "(Q,R)"} where @{term "Q"} is an matrix whose columns are orthogonal unit vectors, @{term "R"} is upper triangular and @{term "A=Q**R"}.\ definition "QR_decomposition A = (let Q = divide_by_norm (Gram_Schmidt_matrix A) in (Q, (transpose Q) ** A))" lemma is_basis_columns_fst_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes b: "is_basis (columns A)" and c: "card (columns A) = ncols A" shows "is_basis (columns (fst (QR_decomposition A))) \ card (columns (fst (QR_decomposition A))) = ncols A" proof (rule conjI, unfold is_basis_def, rule conjI) have "vec.span (columns (fst (QR_decomposition A))) = vec.span (columns (Gram_Schmidt_matrix A))" unfolding vec.span_eq proof (auto) fix x show "x \ vec.span (columns (Gram_Schmidt_matrix A))" using assms(1) assms(2) is_basis_columns_Gram_Schmidt_matrix is_basis_def by auto next fix x assume x: "x \ columns (Gram_Schmidt_matrix A)" from this obtain i where x_col_i: "x=column i (Gram_Schmidt_matrix A)" unfolding columns_def by blast have zero_not_in: "x \ 0" using is_basis_columns_Gram_Schmidt_matrix[OF b c] unfolding is_basis_def using vec.dependent_zero[of "(columns (Gram_Schmidt_matrix A))"] x by auto have "x=column i (Gram_Schmidt_matrix A)" using x_col_i . also have "... = norm (column i (Gram_Schmidt_matrix A)) *\<^sub>R column i (divide_by_norm (Gram_Schmidt_matrix A))" using zero_not_in unfolding x_col_i columns_def divide_by_norm_def column_def normalize_def by vector finally have x_eq: "x = norm (column i (Gram_Schmidt_matrix A)) *\<^sub>R column i (divide_by_norm (Gram_Schmidt_matrix A))" . show "x \ vec.span (columns (fst (QR_decomposition A)))" unfolding x_eq span_vec_eq apply (rule subspace_mul) apply (auto simp add: columns_def QR_decomposition_def Let_def subspace_span intro: span_superset) using span_superset by force qed thus s: "vec.span (columns (fst (QR_decomposition A))) = (UNIV::(real^'m::{mod_type}) set)" using is_basis_columns_Gram_Schmidt_matrix[OF b c] unfolding is_basis_def by simp thus "card (columns (fst (QR_decomposition A))) = ncols A" by (metis (hide_lams, mono_tags) b c card_columns_le_ncols vec.card_le_dim_spanning finite_columns vec.indep_card_eq_dim_span is_basis_def ncols_def top_greatest) thus "vec.independent (columns (fst (QR_decomposition A)))" by (metis s b c vec.card_eq_dim_span_indep finite_columns vec.indep_card_eq_dim_span is_basis_def) qed lemma orthogonal_fst_QR_decomposition: shows "pairwise orthogonal (columns (fst (QR_decomposition A)))" unfolding pairwise_def columns_def proof (auto) fix i ia assume col_not_eq: "column i (fst (QR_decomposition A)) \ column ia (fst (QR_decomposition A))" hence i_not_ia: "i \ ia" by auto from col_not_eq obtain a where "(fst (QR_decomposition A)) $ a $ i \ (fst (QR_decomposition A)) $ a $ ia" unfolding column_def by force hence col_not_eq2: " (column i (Gram_Schmidt_matrix A)) \ (column ia (Gram_Schmidt_matrix A))" using col_not_eq unfolding QR_decomposition_def Let_def fst_conv by (metis (lifting) divide_by_norm_def vec_lambda_beta) have d1: "column i (fst (QR_decomposition A)) = (1 / norm (\ ia. Gram_Schmidt_matrix A $ ia $ i)) *\<^sub>R (column i (Gram_Schmidt_matrix A))" unfolding QR_decomposition_def Let_def fst_conv unfolding divide_by_norm_def column_def normalize_def unfolding vec_eq_iff by auto have d2: "column ia (fst (QR_decomposition A)) = (1 / norm (\ i. Gram_Schmidt_matrix A $ i $ ia)) *\<^sub>R (column ia (Gram_Schmidt_matrix A))" unfolding QR_decomposition_def Let_def fst_conv unfolding divide_by_norm_def column_def normalize_def unfolding vec_eq_iff by auto show "orthogonal (column i (fst (QR_decomposition A))) (column ia (fst (QR_decomposition A)))" unfolding d1 d2 apply (rule orthogonal_mult) using orthogonal_Gram_Schmidt_matrix[of A] unfolding pairwise_def using col_not_eq2 by auto qed lemma qk_uk_norm: "(1/(norm (column k ((Gram_Schmidt_matrix A))))) *\<^sub>R (column k ((Gram_Schmidt_matrix A))) = column k (fst(QR_decomposition A))" unfolding QR_decomposition_def Let_def fst_conv divide_by_norm_def unfolding column_def normalize_def by vector lemma norm_columns_fst_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes "rank A = ncols A" shows "norm (column i (fst (QR_decomposition A))) = 1" proof - have "vec.independent (columns (Gram_Schmidt_matrix A))" by (metis assms full_rank_imp_is_basis2 independent_columns_Gram_Schmidt_matrix) hence "column i (Gram_Schmidt_matrix A) \ 0" using vec.dependent_zero[of "columns (Gram_Schmidt_matrix A)"] unfolding columns_def by auto thus "norm (column i (fst (QR_decomposition A))) = 1" unfolding QR_decomposition_def Let_def fst_conv by (rule norm_column_divide_by_norm) qed corollary span_fst_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" shows "vec.span (columns A) = vec.span (columns (fst (QR_decomposition A)))" unfolding span_Gram_Schmidt_matrix[of A] unfolding QR_decomposition_def Let_def fst_conv by (metis \span (columns A) = span (columns (Gram_Schmidt_matrix A))\ span_columns_divide_by_norm span_vec_eq) corollary col_space_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" shows "col_space A = col_space (fst (QR_decomposition A))" unfolding col_space_def using span_fst_QR_decomposition by auto lemma independent_columns_fst_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes b: "vec.independent (columns A)" and c: "card (columns A) = ncols A" shows "vec.independent (columns (fst (QR_decomposition A))) \ card (columns (fst (QR_decomposition A))) = ncols A" proof - have r: "rank A = ncols A" thm is_basis_imp_full_rank proof - have "rank A = col_rank A" unfolding rank_col_rank .. also have "... = vec.dim (col_space A)" unfolding col_rank_def .. also have "... = card (columns A)" unfolding col_space_def using b by (rule vec.dim_span_eq_card_independent) also have "... = ncols A" using c . finally show ?thesis . qed have "vec.independent (columns (fst (QR_decomposition A)))" by (metis b c col_rank_def col_space_QR_decomposition col_space_def full_rank_imp_is_basis2 vec.indep_card_eq_dim_span ncols_def rank_col_rank) moreover have "card (columns (fst (QR_decomposition A))) = ncols A" by (metis col_space_QR_decomposition full_rank_imp_is_basis2 ncols_def r rank_eq_dim_col_space') ultimately show ?thesis by simp qed lemma orthogonal_matrix_fst_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" shows "transpose (fst (QR_decomposition A)) ** (fst (QR_decomposition A)) = mat 1" proof (unfold vec_eq_iff, clarify, unfold mat_1_fun, auto) define Q where "Q = fst (QR_decomposition A)" have n: "\i. norm (column i Q) = 1" unfolding Q_def using norm_columns_fst_QR_decomposition[OF r] by auto have c: "card (columns Q) = ncols A" unfolding Q_def by (metis full_rank_imp_is_basis2 independent_columns_fst_QR_decomposition r) have p: "pairwise orthogonal (columns Q)" by (metis Q_def orthogonal_fst_QR_decomposition) fix ia have "(transpose Q ** Q) $ ia $ ia = column ia Q \ column ia Q" unfolding matrix_matrix_mult_inner_mult unfolding row_transpose .. also have "... = 1" using n norm_eq_1 by blast finally show "(transpose Q ** Q) $ ia $ ia = 1" . fix i assume i_not_ia: "i \ ia" have column_i_not_ia: "column i Q \ column ia Q" proof (rule ccontr, simp) assume col_i_ia: "column i Q = column ia Q" have rw: "(\i. column i Q)` (UNIV-{ia}) = {column i Q|i. i\ia}" unfolding columns_def by auto have "card (columns Q) = card ({column i Q|i. i\ia})" by (rule bij_betw_same_card[of id], unfold bij_betw_def columns_def, auto, metis col_i_ia i_not_ia) also have "... = card ((\i. column i Q)` (UNIV-{ia}))" unfolding rw .. also have "... \ card (UNIV - {ia})" by (metis card_image_le finite_code) also have "... < CARD ('n)" by simp finally show False using c unfolding ncols_def by simp qed hence oia: "orthogonal (column i Q) (column ia Q)" using p unfolding pairwise_def unfolding columns_def by auto have "(transpose Q ** Q) $ i $ ia = column i Q \ column ia Q" unfolding matrix_matrix_mult_inner_mult unfolding row_transpose .. also have "... = 0" using oia unfolding orthogonal_def . finally show "(transpose Q ** Q) $ i $ ia = 0" . qed corollary orthogonal_matrix_fst_QR_decomposition': fixes A::"real^'n::{mod_type}^'n::{mod_type}" assumes "rank A = ncols A" shows "orthogonal_matrix (fst (QR_decomposition A))" by (metis assms orthogonal_matrix orthogonal_matrix_fst_QR_decomposition) lemma column_eq_fst_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" and c: "column i (fst (QR_decomposition A)) = column ia (fst (QR_decomposition A))" shows "i = ia" proof (rule ccontr) assume i_not_ia: "i \ ia" have "columns (fst (QR_decomposition A)) = (\x. column x (fst (QR_decomposition A)))` (UNIV::('n::{mod_type}) set)" unfolding columns_def by auto also have "... = (\x. column x (fst (QR_decomposition A)))` ((UNIV::('n::{mod_type}) set)-{ia})" proof (unfold image_def, auto) fix xa show "\x\UNIV - {ia}. column xa (fst (QR_decomposition A)) = column x (fst (QR_decomposition A))" proof (cases "xa = ia") case True thus ?thesis using c i_not_ia by (metis DiffI UNIV_I empty_iff insert_iff) next case False thus ?thesis by auto qed qed finally have columns_rw: "columns (fst (QR_decomposition A)) = (\x. column x (fst (QR_decomposition A))) ` (UNIV - {ia})" . have "ncols A = card (columns (fst (QR_decomposition A)))" by (metis full_rank_imp_is_basis2 independent_columns_fst_QR_decomposition r) also have "... \ card (UNIV - {ia})" unfolding columns_rw by (rule card_image_le, simp) also have "... = card (UNIV::'n set) - 1" by (simp add: card_Diff_singleton) finally show False unfolding ncols_def - by (metis Nat.add_0_right Nat.le_diff_conv2 One_nat_def Suc_n_not_le_n add_Suc_right one_le_card_finite) + by (metis Nat.add_0_right le_diff_conv2 One_nat_def Suc_n_not_le_n add_Suc_right one_le_card_finite) qed corollary column_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" shows "column k ((Gram_Schmidt_matrix A)) = (column k A) - (\x\{column i (fst (QR_decomposition A))|i. i < k}. (x \ (column k A) / (x \ x)) *\<^sub>R x)" proof - let ?uk="column k ((Gram_Schmidt_matrix A))" let ?qk="column k (fst(QR_decomposition A))" let ?ak="(column k A)" define f where "f x = (1/norm x) *\<^sub>R x" for x :: "real^'m::{mod_type}" let ?g="\x::real^'m::{mod_type}. (x \ (column k A) / (x \ x)) *\<^sub>R x" have set_rw: "{column i (fst (QR_decomposition A))|i. i < k} = f`{column i (Gram_Schmidt_matrix A)|i. i < k}" proof (auto) fix i assume i: "i < k" have col_rw: "column i (fst (QR_decomposition A)) = (1/norm (column i (Gram_Schmidt_matrix A))) *\<^sub>R (column i (Gram_Schmidt_matrix A))" unfolding QR_decomposition_def Let_def fst_conv divide_by_norm_def column_def normalize_def by vector thus "column i (fst (QR_decomposition A)) \ f ` {column i (Gram_Schmidt_matrix A) |i. i < k}" unfolding f_def using i by auto show "\ia. f (column i (Gram_Schmidt_matrix A)) = column ia (fst (QR_decomposition A)) \ ia < k" by (rule exI[of _ i], simp add: f_def col_rw i) qed have "(\x\{column i (fst (QR_decomposition A))|i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x) = (\x\(f`{column i (Gram_Schmidt_matrix A)|i. i < k}). (x \ ?ak / (x \ x)) *\<^sub>R x)" unfolding set_rw .. also have "... = sum (?g \ f) {column i (Gram_Schmidt_matrix A)|i. i < k}" proof (rule sum.reindex, unfold inj_on_def, auto) fix i ia assume i: "i < k" and ia: "ia < k" and f_eq: "f (column i (Gram_Schmidt_matrix A)) = f (column ia (Gram_Schmidt_matrix A))" have fi: "f (column i (Gram_Schmidt_matrix A)) = column i (fst (QR_decomposition A))" unfolding f_def QR_decomposition_def Let_def fst_conv divide_by_norm_def column_def normalize_def by vector have fia: "f (column ia (Gram_Schmidt_matrix A)) = column ia (fst (QR_decomposition A))" unfolding f_def QR_decomposition_def Let_def fst_conv divide_by_norm_def column_def normalize_def by vector have "i = ia" using column_eq_fst_QR_decomposition[OF r] f_eq unfolding fi fia by simp thus "column i (Gram_Schmidt_matrix A) = column ia (Gram_Schmidt_matrix A)" by simp qed also have "... = (\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. ((1 / norm x) *\<^sub>R x \ ?ak / ((1 / norm x) *\<^sub>R x \ (1 / norm x) *\<^sub>R x)) *\<^sub>R (1 / norm x) *\<^sub>R x)" unfolding o_def f_def .. also have "... = (\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. ((1 / norm x) *\<^sub>R x \ ?ak) *\<^sub>R (1 / norm x) *\<^sub>R x)" proof (rule sum.cong, simp) fix x assume x: "x \ {column i (Gram_Schmidt_matrix A) |i. i < k}" have "vec.independent {column i (Gram_Schmidt_matrix A) |i. i < k}" proof (rule vec.independent_mono[of "columns (Gram_Schmidt_matrix A)"]) show "vec.independent (columns (Gram_Schmidt_matrix A))" using full_rank_imp_is_basis2[of "(Gram_Schmidt_matrix A)"] by (metis full_rank_imp_is_basis2 independent_columns_Gram_Schmidt_matrix r) show "{column i (Gram_Schmidt_matrix A) |i. i < k} \ columns (Gram_Schmidt_matrix A)" unfolding columns_def by auto qed hence "x \ 0" using vec.dependent_zero[of " {column i (Gram_Schmidt_matrix A) |i. i < k}"] x by blast hence "((1 / norm x) *\<^sub>R x \ (1 / norm x) *\<^sub>R x) = 1" by (metis inverse_eq_divide norm_eq_1 norm_sgn sgn_div_norm) thus "((1 / norm x) *\<^sub>R x \ ?ak / ((1 / norm x) *\<^sub>R x \ (1 / norm x) *\<^sub>R x)) *\<^sub>R (1 / norm x) *\<^sub>R x = ((1 / norm x) *\<^sub>R x \ column k A) *\<^sub>R (1 / norm x) *\<^sub>R x" by auto qed also have "... = (\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. (((x \ ?ak)) / (x \ x)) *\<^sub>R x)" proof (rule sum.cong, simp) fix x assume x: "x \ {column i (Gram_Schmidt_matrix A) |i. i < k}" show "((1 / norm x) *\<^sub>R x \ column k A) *\<^sub>R (1 / norm x) *\<^sub>R x = (x \ column k A / (x \ x)) *\<^sub>R x" by (metis (hide_lams, no_types) mult.right_neutral inner_commute inner_scaleR_right norm_cauchy_schwarz_eq scaleR_one scaleR_scaleR times_divide_eq_right times_divide_times_eq) qed finally have "?ak - (\x\{column i (fst (QR_decomposition A))|i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x) = ?ak - (\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. (((x \ ?ak)) / (x \ x)) *\<^sub>R x)" by auto also have "... = ?uk" using column_Gram_Schmidt_matrix[of k A] by auto finally show ?thesis .. qed lemma column_QR_decomposition': fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" shows "(column k A) = column k ((Gram_Schmidt_matrix A)) + (\x\{column i (fst (QR_decomposition A))|i. i < k}. (x \ (column k A) / (x \ x)) *\<^sub>R x)" using column_QR_decomposition[OF r] by simp lemma norm_uk_eq: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" shows "norm (column k ((Gram_Schmidt_matrix A))) = ((column k (fst(QR_decomposition A))) \ (column k A))" proof - let ?uk="column k ((Gram_Schmidt_matrix A))" let ?qk="column k (fst(QR_decomposition A))" let ?ak="(column k A)" have sum_rw: "(?uk \ (\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x)) = 0" proof - have "(?uk \ (\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x)) = ((\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. ?uk \ ((x \ ?ak / (x \ x)) *\<^sub>R x)))" unfolding inner_sum_right .. also have "... = (\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. ((x \ ?ak / (x \ x)) * (?uk \ x)))" unfolding inner_scaleR_right .. also have "... = 0" proof (rule sum.neutral, clarify) fix x i assume "i column i (Gram_Schmidt_matrix A) = 0" by (metis less_irrefl r scaleR_columns_Gram_Schmidt_matrix) thus "column i (Gram_Schmidt_matrix A) \ ?ak / (column i (Gram_Schmidt_matrix A) \ column i (Gram_Schmidt_matrix A)) * (?uk \ column i (Gram_Schmidt_matrix A)) = 0" by auto qed finally show ?thesis . qed have "?qk \ ?ak = ((1/(norm ?uk)) *\<^sub>R ?uk) \ ?ak" unfolding qk_uk_norm .. also have "... = (1/(norm ?uk)) * (?uk \ ?ak)" unfolding inner_scaleR_left .. also have "... = (1/(norm ?uk)) * (?uk \ (?uk + (\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x)))" using column_Gram_Schmidt_matrix2[of k A] by auto also have "... = (1/(norm ?uk)) * ((?uk \ ?uk) + (?uk \ (\x\{column i (Gram_Schmidt_matrix A) |i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x)))" unfolding inner_add_right .. also have "... = (1/(norm ?uk)) * (?uk \ ?uk)" unfolding sum_rw by auto also have "... = norm ?uk" by (metis abs_of_nonneg divide_eq_imp div_by_0 inner_commute inner_ge_zero inner_real_def norm_mult_vec real_inner_1_right real_norm_def times_divide_eq_right) finally show ?thesis .. qed corollary column_QR_decomposition2: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" shows "(column k A) = (\x\{column i (fst (QR_decomposition A))|i. i \ k}. (x \ (column k A)) *\<^sub>R x)" proof - let ?uk="column k ((Gram_Schmidt_matrix A))" let ?qk="column k (fst(QR_decomposition A))" let ?ak="(column k A)" have set_rw: "{column i (fst (QR_decomposition A))|i. i \ k} = insert (column k (fst (QR_decomposition A))) {column i (fst (QR_decomposition A))|i. i < k}" by (auto, metis less_linear not_less) have uk_norm_uk_qk: "?uk = norm ?uk *\<^sub>R ?qk" proof - have "vec.independent (columns (Gram_Schmidt_matrix A))" by (metis full_rank_imp_is_basis2 independent_columns_Gram_Schmidt_matrix r) moreover have "?uk \ columns (Gram_Schmidt_matrix A)" unfolding columns_def by auto ultimately have "?uk \ 0" using vec.dependent_zero[of "columns (Gram_Schmidt_matrix A)"] unfolding columns_def by auto hence norm_not_0: "norm ?uk \ 0" unfolding norm_eq_zero . have "norm (?uk) *\<^sub>R ?qk = (norm ?uk) *\<^sub>R ((1 / norm ?uk) *\<^sub>R ?uk)" using qk_uk_norm[of k A] by simp also have "... = ((norm ?uk) * (1 / norm ?uk)) *\<^sub>R ?uk" unfolding scaleR_scaleR .. also have "... = ?uk" using norm_not_0 by auto finally show ?thesis .. qed have norm_qk_1: "?qk \ ?qk = 1" using norm_eq_1 norm_columns_fst_QR_decomposition[OF r] by auto have "?ak = ?uk + (\x\{column i (fst (QR_decomposition A))|i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x)" using column_QR_decomposition'[OF r] by auto also have "... = (norm ?uk *\<^sub>R ?qk) + (\x\{column i (fst (QR_decomposition A))|i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x)" using uk_norm_uk_qk by simp also have "... = ((?qk \ ?ak) *\<^sub>R ?qk) + (\x\{column i (fst (QR_decomposition A))|i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x)" unfolding norm_uk_eq[OF r] .. also have "... = ((?qk \ ?ak)/(?qk \ ?qk)) *\<^sub>R ?qk + (\x\{column i (fst (QR_decomposition A))|i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x)" using norm_qk_1 by fastforce also have "... = (\x\insert ?qk {column i (fst (QR_decomposition A))|i. i < k}. (x \ ?ak / (x \ x)) *\<^sub>R x)" proof (rule sum.insert[symmetric]) show "finite {column i (fst (QR_decomposition A)) |i. i < k}" by simp show "column k (fst (QR_decomposition A)) \ {column i (fst (QR_decomposition A)) |i. i < k}" proof (rule ccontr, simp) assume "\i. column k (fst (QR_decomposition A)) = column i (fst (QR_decomposition A)) \ i < k" from this obtain i where col_eq: "column k (fst (QR_decomposition A)) = column i (fst (QR_decomposition A))" and i_less_k: "i < k" by blast show False using column_eq_fst_QR_decomposition[OF r col_eq] i_less_k by simp qed qed also have "... = (\x\{column i (fst (QR_decomposition A))|i. i \ k}. (x \ (column k A)) *\<^sub>R x)" proof (rule sum.cong, simp add: set_rw) fix x assume x: "x \ {column i (fst (QR_decomposition A)) |i. i \ k}" from this obtain i where i: "x=column i (fst (QR_decomposition A))" by blast hence "(x \ x) = 1" using norm_eq_1 norm_columns_fst_QR_decomposition[OF r] by auto thus "(x \ column k A / (x \ x)) *\<^sub>R x = (x \ column k A) *\<^sub>R x" by simp qed finally show ?thesis . qed lemma orthogonal_columns_fst_QR_decomposition: assumes i_not_ia: "(column i (fst (QR_decomposition A))) \ (column ia (fst (QR_decomposition A)))" shows "(column i (fst (QR_decomposition A)) \ column ia (fst (QR_decomposition A))) = 0" proof - have i: "column i (fst (QR_decomposition A)) \ columns (fst (QR_decomposition A))" unfolding columns_def by auto have ia: "column ia (fst (QR_decomposition A)) \ columns (fst (QR_decomposition A))" unfolding columns_def by auto show ?thesis using orthogonal_fst_QR_decomposition[of A] i ia i_not_ia unfolding pairwise_def orthogonal_def by auto qed lemma scaler_column_fst_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes i: "i>j" and r: "rank A = ncols A" shows "column i (fst (QR_decomposition A)) \ column j A = 0" proof - have "column i (fst(QR_decomposition A)) \ column j A = column i (fst (QR_decomposition A)) \ (\x\{column i (fst (QR_decomposition A))|i. i \ j}. (x \ (column j A)) *\<^sub>R x)" using column_QR_decomposition2[OF r] by presburger also have "... = (\x\{column i (fst (QR_decomposition A))|i. i \ j}. column i (fst (QR_decomposition A)) \ (x \ (column j A)) *\<^sub>R x)" unfolding real_inner_class.inner_sum_right .. also have "... = (\x\{column i (fst (QR_decomposition A))|i. i \ j}. (x \ (column j A)) *(column i (fst (QR_decomposition A)) \ x))" unfolding real_inner_class.inner_scaleR_right .. also have "... = 0" proof (rule sum.neutral, clarify) fix ia assume ia: "ia \ j" have i_not_ia: "i \ ia" using i ia by simp hence "(column i (fst (QR_decomposition A)) \ column ia (fst (QR_decomposition A)))" by (metis column_eq_fst_QR_decomposition r) hence "(column i (fst (QR_decomposition A)) \ column ia (fst (QR_decomposition A))) = 0" by (rule orthogonal_columns_fst_QR_decomposition) thus "column ia (fst (QR_decomposition A)) \ column j A * (column i (fst (QR_decomposition A)) \ column ia (fst (QR_decomposition A))) = 0" by auto qed finally show ?thesis . qed lemma R_Qi_Aj: fixes A::"real^'n::{mod_type}^'m::{mod_type}" shows "(snd (QR_decomposition A)) $ i $ j = column i (fst (QR_decomposition A)) \ column j A" unfolding QR_decomposition_def Let_def snd_conv matrix_matrix_mult_inner_mult unfolding row_transpose by auto lemma sums_columns_Q_0: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" shows "(\x\{column i (fst (QR_decomposition A)) |i. i>b}. x \ column b A * x $ a) = 0" proof (rule sum.neutral, auto) fix i assume "b column b A = 0" by (rule scaler_column_fst_QR_decomposition, simp add: r) qed lemma QR_decomposition_mult: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" shows "A = (fst (QR_decomposition A)) ** (snd (QR_decomposition A))" proof - have "\b. column b A = column b ((fst (QR_decomposition A)) ** (snd (QR_decomposition A)))" proof (clarify) fix b have "(fst (QR_decomposition A) ** snd (QR_decomposition A)) = (\ i j. \k\UNIV. fst (QR_decomposition A) $ i $ k * (column k (fst (QR_decomposition A)) \ column j A))" unfolding matrix_matrix_mult_def R_Qi_Aj by auto hence "column b ((fst (QR_decomposition A) ** snd (QR_decomposition A))) = column b ((\ i j. \k\UNIV. fst (QR_decomposition A) $ i $ k * (column k (fst (QR_decomposition A)) \ column j A)))" by auto also have "... = (\x\{column i (fst (QR_decomposition A)) |i. i \ b}. (x \ column b A) *\<^sub>R x)" proof (subst column_def, subst vec_eq_iff, auto) fix a define f where "f i = column i (fst (QR_decomposition A))" for i define g where "g x = (THE i. x = column i (fst (QR_decomposition A)))" for x have f_eq: "f`UNIV = {column i (fst (QR_decomposition A)) |i. i\UNIV}" unfolding f_def by auto have inj_f: "inj f" by (metis inj_on_def f_def column_eq_fst_QR_decomposition r) have "(\x\{column i (fst (QR_decomposition A)) |i. i \ b}. x \ column b A * x $ a) = (\x\{column i (fst (QR_decomposition A)) |i. i\UNIV}. x \ column b A * x $ a)" proof - let ?c= "{column i (fst (QR_decomposition A)) |i. i\UNIV}" let ?d= "{column i (fst (QR_decomposition A)) |i. i\b}" let ?f = "{column i (fst (QR_decomposition A)) |i. i>b}" have set_rw: "?c = ?d \ ?f" by force have "(\x\?c. x \ column b A * x $ a) = (\x\(?d \ ?f). x \ column b A * x $ a)" using set_rw by simp also have "... = (\x\?d. x \ column b A * x $ a) + (\x\?f. x \ column b A * x $ a)" by (rule sum.union_disjoint, auto, metis f_def inj_eq inj_f not_le) also have "... = (\x\?d. x \ column b A * x $ a)" using sums_columns_Q_0[OF r] by auto finally show ?thesis .. qed also have "... = (\x\f`UNIV. x \ column b A * x $ a)" using f_eq by auto also have "... = (\k\UNIV. fst (QR_decomposition A) $ a $ k * (column k (fst (QR_decomposition A)) \ column b A))" unfolding sum.reindex[OF inj_f] unfolding f_def column_def by (rule sum.cong, simp_all) finally show " (\k\UNIV. fst (QR_decomposition A) $ a $ k * (column k (fst (QR_decomposition A)) \ column b A)) = (\x\{column i (fst (QR_decomposition A)) |i. i \ b}. x \ column b A * x $ a)" .. qed also have "... = column b A" using column_QR_decomposition2[OF r] by simp finally show "column b A = column b (fst (QR_decomposition A) ** snd (QR_decomposition A))" .. qed thus ?thesis unfolding column_def vec_eq_iff by auto qed lemma upper_triangular_snd_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" shows "upper_triangular (snd (QR_decomposition A))" proof (unfold upper_triangular_def, auto) fix i j::'n assume j_less_i: "j < i" have "snd (QR_decomposition A) $ i $ j = column i (fst (QR_decomposition A)) \ column j A" unfolding QR_decomposition_def Let_def fst_conv snd_conv unfolding matrix_matrix_mult_inner_mult row_transpose .. also have "... = 0" using scaler_column_fst_QR_decomposition[OF j_less_i r] . finally show "snd (QR_decomposition A) $ i $ j = 0" by auto qed lemma upper_triangular_invertible: fixes A :: "real^'n::{finite,wellorder}^'n::{finite,wellorder}" assumes u: "upper_triangular A" and d: "\i. A $ i $ i \ 0" shows "invertible A" proof - have det_R: "det A = (prod (\i. A$i$i) (UNIV::'n set))" using det_upperdiagonal u unfolding upper_triangular_def by blast also have "... \ 0" using d by auto finally show ?thesis by (metis invertible_det_nz) qed lemma invertible_snd_QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" shows "invertible (snd (QR_decomposition A))" proof (rule upper_triangular_invertible) show "upper_triangular (snd (QR_decomposition A))" using upper_triangular_snd_QR_decomposition[OF r] . show "\i. snd (QR_decomposition A) $ i $ i \ 0" proof (rule allI) fix i have ind: "vec.independent (columns (Gram_Schmidt_matrix A))" by (metis full_rank_imp_is_basis2 independent_columns_Gram_Schmidt_matrix r) hence zero_not_in: "0 \ (columns (Gram_Schmidt_matrix A))" by (metis vec.dependent_zero) hence c:"column i (Gram_Schmidt_matrix A) \ 0" unfolding columns_def by simp have "snd (QR_decomposition A) $ i $ i = column i (fst (QR_decomposition A)) \ column i A" unfolding QR_decomposition_def Let_def snd_conv fst_conv unfolding matrix_matrix_mult_inner_mult unfolding row_transpose .. also have "... = norm (column i (Gram_Schmidt_matrix A))" unfolding norm_uk_eq[OF r, symmetric] .. also have "... \ 0" by (rule ccontr, simp add: c) finally show "snd (QR_decomposition A) $ i $ i \ 0" . qed qed lemma QR_decomposition: fixes A::"real^'n::{mod_type}^'m::{mod_type}" assumes r: "rank A = ncols A" shows "A = fst (QR_decomposition A) ** snd (QR_decomposition A) \ pairwise orthogonal (columns (fst (QR_decomposition A))) \ (\i. norm (column i (fst (QR_decomposition A))) = 1) \ (transpose (fst (QR_decomposition A))) ** (fst (QR_decomposition A)) = mat 1 \ vec.independent (columns (fst (QR_decomposition A))) \ col_space A = col_space (fst (QR_decomposition A)) \ card (columns A) = card (columns (fst (QR_decomposition A))) \ invertible (snd (QR_decomposition A)) \ upper_triangular (snd (QR_decomposition A))" by (metis QR_decomposition_mult col_space_def full_rank_imp_is_basis2 independent_columns_fst_QR_decomposition invertible_snd_QR_decomposition norm_columns_fst_QR_decomposition orthogonal_fst_QR_decomposition orthogonal_matrix_fst_QR_decomposition r span_fst_QR_decomposition upper_triangular_snd_QR_decomposition) lemma QR_decomposition_square: fixes A::"real^'n::{mod_type}^'n::{mod_type}" assumes r: "rank A = ncols A" shows "A = fst (QR_decomposition A) ** snd (QR_decomposition A) \ orthogonal_matrix (fst (QR_decomposition A)) \ upper_triangular (snd (QR_decomposition A)) \ invertible (snd (QR_decomposition A)) \ pairwise orthogonal (columns (fst (QR_decomposition A))) \ (\i. norm (column i (fst (QR_decomposition A))) = 1) \ vec.independent (columns (fst (QR_decomposition A))) \ col_space A = col_space (fst (QR_decomposition A)) \ card (columns A) = card (columns (fst (QR_decomposition A)))" by (metis QR_decomposition orthogonal_matrix_fst_QR_decomposition' r) text\QR for computing determinants\ lemma det_QR_decomposition: fixes A::"real^'n::{mod_type}^'n::{mod_type}" assumes r: "rank A = ncols A" shows "\det A\ = \(prod (\i. snd(QR_decomposition A)$i$i) (UNIV::'n set))\" proof - let ?Q="fst(QR_decomposition A)" let ?R="snd(QR_decomposition A)" have det_R: "det ?R = (prod (\i. snd(QR_decomposition A)$i$i) (UNIV::'n set))" apply (rule det_upperdiagonal) using upper_triangular_snd_QR_decomposition[OF r] unfolding upper_triangular_def by simp have "\det A\ = \det ?Q * det ?R\" by (metis QR_decomposition_mult det_mul r) also have "... = \det ?Q\ * \det ?R\" unfolding abs_mult .. also have "... = 1 * \det ?R\" using det_orthogonal_matrix[OF orthogonal_matrix_fst_QR_decomposition'[OF r]] by auto also have "... = \det ?R\" by simp also have "... = \(prod (\i. snd(QR_decomposition A)$i$i) (UNIV::'n set))\" unfolding det_R .. finally show ?thesis . qed end diff --git a/thys/TortoiseHare/Brent.thy b/thys/TortoiseHare/Brent.thy --- a/thys/TortoiseHare/Brent.thy +++ b/thys/TortoiseHare/Brent.thy @@ -1,232 +1,232 @@ (*<*) theory Brent imports Basis begin (*>*) section\ Brent's algorithm \label{sec:brent} \ text\ @{cite "Brent:1980"} improved on the Tortoise and Hare algorithm and used it to factor large primes. In practice it makes significantly fewer calls to the function \f\ before detecting a loop. We begin by defining the base-2 logarithm. \ fun lg :: "nat \ nat" where [simp del]: "lg x = (if x \ 1 then 0 else 1 + lg (x div 2))" lemma lg_safe: "lg 0 = 0" "lg (Suc 0) = 0" "lg (Suc (Suc 0)) = 1" "0 < x \ lg (x + x) = 1 + lg x" by (simp_all add: lg.simps) lemma lg_inv: "0 < x \ lg (2 ^ x) = x" proof(induct x) case (Suc x) then show ?case by (cases x, simp_all add: lg.simps Suc_lessI not_le) qed simp lemma lg_inv2: "2 ^ i = x \ 2 ^ lg x = x" by (induct i arbitrary: x) (auto simp: lg.simps, arith) lemmas lg_simps = lg_safe lg_inv lg_inv2 subsection\ Finding \lambda\ \ text (in properties) \ Imagine now that the Tortoise carries an unbounded number of carrots, which he passes to the Hare when they meet, and the Hare has a teleporter. The Hare eats a carrot each time she waits for the function @{term "f"} to execute, and initially has just one. If she runs out of carrots before meeting the Tortoise again, she teleports him to her position, and he gives her twice as many carrots as the last time they met (tracked by the variable \carrots\). By counting how many carrots she has eaten from when she last teleported the Tortoise (recorded in \l\) until she finally has surplus carrots when she meets him again, the Hare directly discovers @{term "lambda"}. \ record 'a state = m :: nat \ \\\\\ l :: nat \ \\\\\ carrots :: nat hare :: "'a" tortoise :: "'a" context properties begin definition (in fx0) find_lambda :: "'a state \ 'a state" where "find_lambda \ (\s. s\ carrots := 1, l := 1, tortoise := x0, hare := f x0 \) ;; while (hare \<^bold>\ tortoise) ( ( \<^bold>if carrots \<^bold>= l \<^bold>then (\s. s\ tortoise := hare s, carrots := 2 * carrots s, l := 0 \) \<^bold>else SKIP ) ;; (\s. s\ hare := f (hare s), l := l s + 1 \) )" text\ The termination argument goes intuitively as follows. The Hare eats as many carrots as it takes to teleport the Tortoise into the loop. Afterwards she continues the teleportation dance until the Tortoise has given her enough carrots to make it all the way around the loop and back to him. We can calculate the Tortoise's position as a function of \carrots\. \ definition carrots_total :: "nat \ nat" where "carrots_total c \ \i carrots_total (c + c) = c + carrots_total c" by (auto simp: carrots_total_def lg_simps) definition find_lambda_measures :: "( (nat \ nat) \ (nat \ nat) ) set" where "find_lambda_measures \ measures [\(l, c). mu - carrots_total c, \(l, c). LEAST i. lambda \ c * 2^i, \(l, c). c - l]" lemma find_lambda_measures_wellfounded: "wf find_lambda_measures" by (simp add: find_lambda_measures_def) lemma find_lambda_measures_decreases1: assumes "c = 2 ^ i" assumes "mu \ carrots_total c \ c \ lambda" assumes "seq (carrots_total c) \ seq (carrots_total c + c)" shows "( (c', 2 * c), (c, c) ) \ find_lambda_measures" proof(cases "mu \ carrots_total c") case False with assms show ?thesis by (auto simp: find_lambda_measures_def carrots_total_simps mult_2 field_simps diff_less_mono2) next case True { fix x assume x: "(0::nat) < x" have "\n. lambda \ x * 2 ^ n" proof(induct lambda) case (Suc i) then obtain n where "i \ x * 2 ^ n" by blast with x show ?case by (clarsimp intro!: exI[where x="Suc n"] simp: field_simps mult_2) - (metis Nat.add_0_right Suc_leI linorder_neqE_nat mult_eq_0_iff nat_add_left_cancel not_le numeral_2_eq_2 old.nat.distinct(2) power_not_zero trans_le_add2) + (metis Nat.add_0_right Suc_leI linorder_neqE_nat mult_eq_0_iff add_left_cancel not_le numeral_2_eq_2 old.nat.distinct(2) power_not_zero trans_le_add2) qed simp } note ex = this have "(LEAST j. lambda \ 2 ^ (i + 1) * 2 ^ j) < (LEAST j. lambda \ 2 ^ i * 2 ^ j)" proof(rule LeastI2_wellorder_ex[OF ex, rotated], rule LeastI2_wellorder_ex[OF ex, rotated]) fix x y assume "lambda \ 2 ^ i * 2 ^ y" "lambda \ 2 ^ (i + 1) * 2 ^ x" "\z. lambda \ 2 ^ (i + 1) * 2 ^ z \ x \ z" with True assms properties_loop[where i="carrots_total c" and j=1] show "x < y" by (cases y, auto simp: less_Suc_eq_le) qed simp_all with True \c = 2 ^ i\ show ?thesis by (clarsimp simp: find_lambda_measures_def mult_2 carrots_total_simps field_simps power_add) qed lemma find_lambda_measures_decreases2: assumes "ls < c" shows "( (Suc ls, c), (ls, c) ) \ find_lambda_measures" using assms by (simp add: find_lambda_measures_def) lemma find_lambda: "\\True\\ find_lambda \l \<^bold>= \lambda\\" apply (simp add: find_lambda_def) apply (rule hoare_pre) apply (rule whileI[where I="\0\ \<^bold>< l \<^bold>\ l \<^bold>\ carrots \<^bold>\ (\mu\ \<^bold>\ carrots_total \ carrots \<^bold>\ l \<^bold>\ \lambda\) \<^bold>\ (\<^bold>\i. carrots \<^bold>= \2^i\) \<^bold>\ tortoise \<^bold>= seq \ carrots_total \ carrots \<^bold>\ hare \<^bold>= seq \ (l \<^bold>+ (carrots_total \ carrots))" and r="inv_image find_lambda_measures (l \<^bold>\ carrots)"] wp_intro)+ using properties_lambda_gt_0 apply (clarsimp simp: field_simps mult_2_right carrots_total_simps) apply (intro conjI impI) apply (metis mult_2 power_Suc) apply (case_tac "mu \ carrots_total (l s)") apply (cut_tac i="carrots_total (l s)" and j="l s" in properties_distinct_contrapos, simp_all add: field_simps)[1] apply (cut_tac i="carrots_total (l s)" and j="l s" in properties_loops_ge_mu, simp_all add: field_simps)[1] apply (cut_tac i="carrots_total (2 ^ x)" and j=1 in properties_loop, simp) apply (fastforce simp: le_eq_less_or_eq field_simps) apply (cut_tac i="carrots_total (2 ^ x)" and j="l s" in properties_loops_ge_mu, simp_all add: field_simps)[1] apply (cut_tac i="carrots_total (2 ^ x)" and j="l s" in properties_distinct_contrapos, simp_all add: field_simps)[1] apply (simp add: find_lambda_measures_wellfounded) apply (clarsimp simp: add.commute find_lambda_measures_decreases1 find_lambda_measures_decreases2) apply (rule wp_intro) using properties_lambda_gt_0 apply (simp add: carrots_total_simps exI[where x=0]) done subsection\ Finding \mu\ \ text\ With @{term "lambda"} in hand, we can find \mu\ using the same approach as for the Tortoise and Hare (\S\ref{sec:th-finding-mu}), after we first move the Hare to @{term "lambda"}. \ definition (in fx0) find_mu :: "'a state \ 'a state" where "find_mu \ (\s. s\ m := 0, tortoise := x0, hare := seq (l s) \) ;; while (hare \<^bold>\ tortoise) (\s. s\ tortoise := f (tortoise s), hare := f (hare s), m := m s + 1 \)" lemma find_mu: "\l \<^bold>= \lambda\\ find_mu \l \<^bold>= \lambda\ \<^bold>\ m \<^bold>= \mu\\" apply (simp add: find_mu_def) apply (rule hoare_pre) apply (rule whileI[where I="l \<^bold>= \lambda\ \<^bold>\ m \<^bold>\ \mu\ \<^bold>\ tortoise \<^bold>= seq \ m \<^bold>\ hare \<^bold>= seq \ (m \<^bold>+ l)" and r="measure (\mu\ \<^bold>- m)"] wp_intro)+ using properties_lambda_gt_0 properties_loop[where i=mu and j=1] apply (fastforce simp: le_less dest: properties_loops_ge_mu) apply simp using properties_loop[where i=mu and j=1, simplified] apply (fastforce simp: le_eq_less_or_eq) apply (rule wp_intro) apply simp done subsection\ Top level \ definition (in fx0) brent :: "'a state \ 'a state" where "brent \ find_lambda ;; find_mu" theorem brent: "\\True\\ brent \l \<^bold>= \lambda\ \<^bold>\ m \<^bold>= \mu\\" unfolding brent_def by (rule find_lambda find_mu wp_intro)+ end corollary brent_correct: assumes s': "s' = fx0.brent f x arbitrary" shows "fx0.properties f x (l s') (m s')" using assms properties.brent[where f=f and ?x0.0=x] by (fastforce intro: fx0.properties_existence[where f=f and ?x0.0=x] simp: Basis.properties_def valid_def) schematic_goal brent_code[code]: "fx0.brent f x = ?code" unfolding fx0.brent_def fx0.find_lambda_def fx0.find_mu_def fcomp_assoc[symmetric] fcomp_comp by (rule refl) export_code fx0.brent in SML (*<*) end (*>*)