diff -r 029bd37d010a -r 0aeb0ea71ef1 Nominal/FSet.thy --- a/Nominal/FSet.thy Wed May 26 16:28:35 2010 +0200 +++ b/Nominal/FSet.thy Wed May 26 16:56:38 2010 +0200 @@ -164,6 +164,10 @@ shows "(op \ ===> op \ ===> op \) op @ op @" by auto +lemma append_rsp_unfolded: + "\x \ y; v \ w\ \ x @ v \ y @ w" + by auto + lemma [quot_respect]: shows "(op \ ===> op \ ===> op =) sub_list sub_list" by (auto simp add: sub_list_def) @@ -373,6 +377,31 @@ then show "concat a \ concat b" by simp qed + + +lemma concat_rsp_unfolded: + "\list_rel op \ a ba; ba \ bb; list_rel op \ bb b\ \ concat a \ concat b" +proof - + fix a b ba bb + assume a: "list_rel op \ a ba" + assume b: "ba \ bb" + assume c: "list_rel op \ bb b" + have "\x. (\xa\set a. x \ set xa) = (\xa\set b. x \ set xa)" proof + fix x + show "(\xa\set a. x \ set xa) = (\xa\set b. x \ set xa)" proof + assume d: "\xa\set a. x \ set xa" + show "\xa\set b. x \ set xa" by (rule concat_rsp_pre[OF a b c d]) + next + assume e: "\xa\set b. x \ set xa" + have a': "list_rel op \ ba a" by (rule list_rel_symp[OF list_eq_equivp, OF a]) + have b': "bb \ ba" by (rule equivp_symp[OF list_eq_equivp, OF b]) + have c': "list_rel op \ b bb" by (rule list_rel_symp[OF list_eq_equivp, OF c]) + show "\xa\set a. x \ set xa" by (rule concat_rsp_pre[OF c' b' a' e]) + qed + qed + then show "concat a \ concat b" by simp +qed + lemma [quot_respect]: "((op =) ===> op \ ===> op \) filter filter" by auto