diff -r 000000000000 -r 1378b654acde thys/Hoare_gen.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys/Hoare_gen.thy Thu Mar 06 13:28:38 2014 +0000 @@ -0,0 +1,602 @@ +header {* + Generic Separation Logic +*} + +theory Hoare_gen +imports Main + (*"../progtut/Tactical" *) + "../Separation_Algebra/Sep_Tactics" +begin + +ML_file "../Separation_Algebra/sep_tactics.ML" + +definition pasrt :: "bool \ (('a::sep_algebra) \ bool)" ("<_>" [72] 71) +where "pasrt b = (\ s . s = 0 \ b)" + +lemma sep_conj_cond1: "(p \* \* q) = ( \* p \* q)" + by(simp add: sep_conj_ac) + +lemma sep_conj_cond2: "(p \* ) = ( \* p)" + by(simp add: sep_conj_ac) + +lemma sep_conj_cond3: "(( \* p) \* r) = ( \* p \* r)" + by (metis sep.mult_assoc) + +lemmas sep_conj_cond = sep_conj_cond1 sep_conj_cond2 sep_conj_cond3 + +lemma cond_true_eq[simp]: " = \" + by(unfold sep_empty_def pasrt_def, auto) + +lemma cond_true_eq1: "( \* p) = p" + by(simp) + +lemma false_simp [simp]: " = sep_false" (* move *) + by (simp add:pasrt_def) + +lemma cond_true_eq2: "(p \* ) = p" + by simp + +lemma condD: "( ** r) s \ b \ r s" +by (unfold sep_conj_def pasrt_def, auto) + +locale sep_exec = + fixes step :: "'conf \ 'conf" + and recse:: "'conf \ 'a::sep_algebra" +begin + +definition "run n = step ^^ n" + +lemma run_add: "run (n1 + n2) s = run n1 (run n2 s)" + apply (unfold run_def) + by (metis funpow_add o_apply) + +definition + Hoare_gen :: "('a \ bool) \ ('a \ bool) \ ('a \ bool) \ bool" + ("(\(1_)\ / (_)/ \(1_)\)" 50) +where + "\ p \ c \ q \ \ + (\ s r. (p**c**r) (recse s) \ (\ k. ((q ** c ** r) (recse (run (Suc k) s)))))" + +lemma HoareI [case_names Pre]: + assumes h: "\ r s. (p**c**r) (recse s) \ (\ k. ((q ** c ** r) (recse (run (Suc k) s))))" + shows "\ p \ c \ q \" + using h + by (unfold Hoare_gen_def, auto) + +lemma frame_rule: + assumes h: "\ p \ c \ q \" + shows "\ p ** r \ c \ q ** r \" +proof(induct rule: HoareI) + case (Pre r' s') + hence "(p \* c \* r \* r') (recse s')" by (auto simp:sep_conj_ac) + from h[unfolded Hoare_gen_def, rule_format, OF this] + show ?case + by (metis sep_conj_assoc sep_conj_left_commute) +qed + +lemma sequencing: + assumes h1: "\p\ c \q\" + and h2: "\q\ c \r\" + shows "\p\ c \r\" +proof(induct rule:HoareI) + case (Pre r' s') + from h1[unfolded Hoare_gen_def, rule_format, OF Pre] + obtain k1 where "(q \* c \* r') (recse (run (Suc k1) s'))" by auto + from h2[unfolded Hoare_gen_def, rule_format, OF this] + obtain k2 where "(r \* c \* r') (recse (run (Suc k2) (run (Suc k1) s')))" by auto + thus ?case + apply (rule_tac x = "Suc (k1 + k2)" in exI) + by (metis add_Suc_right nat_add_commute sep_exec.run_add) +qed + +lemma pre_stren: + assumes h1: "\p\ c \q\" + and h2: "\s. r s \ p s" + shows "\r\ c \q\" +proof(induct rule:HoareI) + case (Pre r' s') + with h2 + have "(p \* c \* r') (recse s')" + by (metis sep_conj_impl1) + from h1[unfolded Hoare_gen_def, rule_format, OF this] + show ?case . +qed + +lemma post_weaken: + assumes h1: "\p\ c \q\" + and h2: "\ s. q s \ r s" + shows "\p\ c \r\" +proof(induct rule:HoareI) + case (Pre r' s') + from h1[unfolded Hoare_gen_def, rule_format, OF this] + obtain k where "(q \* c \* r') (recse (run (Suc k) s'))" by blast + with h2 + show ?case + by (metis sep_conj_impl1) +qed + +lemma hoare_adjust: + assumes h1: "\p1\ c \q1\" + and h2: "\s. p s \ p1 s" + and h3: "\s. q1 s \ q s" + shows "\p\ c \q\" + using h1 h2 h3 post_weaken pre_stren + by (metis) + +lemma code_exI: + assumes h: "\ k. \p\ c(k) \q\" + shows "\p\ EXS k. c(k) \q\" +proof(unfold pred_ex_def, induct rule:HoareI) + case (Pre r' s') + then obtain k where "(p \* (\ s. c k s) \* r') (recse s')" + by (auto elim!:sep_conjE intro!:sep_conjI) + from h[unfolded Hoare_gen_def, rule_format, OF this] + show ?case + by (auto elim!:sep_conjE intro!:sep_conjI) +qed + +lemma code_extension: + assumes h: "\ p \ c \ q \" + shows "\ p \ c ** e \ q \" +proof(induct rule:HoareI) + case (Pre r' s') + hence "(p \* c \* e \* r') (recse s')" by (auto simp:sep_conj_ac) + from h[unfolded Hoare_gen_def, rule_format, OF this] + show ?case + by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac) +qed + +lemma code_extension1: + assumes h: "\ p \ c \ q \" + shows "\ p \ e ** c \ q \" + by (metis code_extension h sep.mult_commute) + +lemma composition: + assumes h1: "\p\ c1 \q\" + and h2: "\q\ c2 \r\" + shows "\p\ c1 ** c2 \r\" +proof(induct rule:HoareI) + case (Pre r' s') + hence "(p \* c1 \* c2 \* r') (recse s')" by (auto simp:sep_conj_ac) + from h1[unfolded Hoare_gen_def, rule_format, OF this] + obtain k1 where "(q \* c2 \* c1 \* r') (recse (run (Suc k1) s'))" + by (auto simp:sep_conj_ac) + from h2[unfolded Hoare_gen_def, rule_format, OF this, folded run_add] + show ?case + by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac) +qed + + +definition + IHoare :: "('b::sep_algebra \ 'a \ bool) \ + ('b \ bool) \ ('a \ bool) \ ('b \ bool) \ bool" + ("(1_).(\(1_)\ / (_)/ \(1_)\)" 50) +where + "I. \P\ c \Q\ = (\s'. \ EXS s.

\* <(s ## s')> \* I(s + s')\ c + \ EXS s. \* <(s ## s')> \* I(s + s')\)" + +lemma IHoareI [case_names IPre]: + assumes h: "\s' s r cnf . (

\* <(s ## s')> \* I(s + s') \* c \* r) (recse cnf) \ + (\k t. ( \* <(t ## s')> \* I(t + s') \* c \* r) (recse (run (Suc k) cnf)))" + shows "I. \P\ c \Q\" + unfolding IHoare_def +proof + fix s' + show " \EXS s.

\* <(s ## s')> \* I (s + s')\ c + \EXS s. \* <(s ## s')> \* I (s + s')\" + proof(unfold pred_ex_def, induct rule:HoareI) + case (Pre r s) + then obtain sa where "(

\* <(sa ## s')> \* I (sa + s') \* c \* r) (recse s)" + by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac) + hence " (\k t. ( \* <(t##s')> \* I(t + s') \* c \* r) (recse (run (Suc k) s)))" + by (rule h) + then obtain k t where h2: "( \* <(t ## s')> \* I(t + s') \* c \* r) (recse (run (Suc k) s))" + by auto + thus "\k. ((\s. \sa. ( \* <(sa ## s')> \* I (sa + s')) s) \* c \* r) (recse (run (Suc k) s))" + apply (rule_tac x = k in exI) + by (auto intro!:sep_conjI elim!:sep_conjE simp:sep_conj_ac) + qed + qed + +lemma I_frame_rule: + assumes h: "I. \P\ c \Q\" + shows "I. \P \* R\ c \Q \* R\" +proof(induct rule:IHoareI) + case (IPre s' s r cnf) + hence "((<(P \* R) s> \* <(s##s')> \* I (s + s')) \* c \* r) (recse cnf)" + by (auto simp:sep_conj_ac) + then obtain s1 s2 + where h1: "((

\* <((s1 + s2) ## s')> \*I (s1 + s2 + s')) \* c \* r) (recse cnf)" + "s1 ## s2" "s1 + s2 ## s'" "P s1" "R s2" + by (unfold pasrt_def, auto elim!:sep_conjE intro!:sep_conjI) + hence "((EXS s.

\* <(s ## s2 +s')> \*I (s + (s2 + s'))) \* c \* r) (recse cnf)" + apply (sep_cancel, unfold pred_ex_def, auto intro!:sep_conjI sep_disj_addI3 elim!:sep_conjE) + apply (rule_tac x = s1 in exI, unfold pasrt_def, + auto intro!:sep_conjI sep_disj_addI3 elim!:sep_conjE simp:sep_conj_ac) + by (metis sep_add_assoc sep_add_disjD) + from h[unfolded IHoare_def Hoare_gen_def, rule_format, OF this] + obtain k s where + "(( \* <(s ## s2 + s')> \* I (s + (s2 + s'))) \* c \* r) (recse (run (Suc k) cnf))" + by (unfold pasrt_def pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI) + thus ?case + proof(rule_tac x = k in exI, rule_tac x = "s + s2" in exI, sep_cancel+) + fix h ha + assume hh: "( \* <(s ## s2 + s')> \* I (s + (s2 + s'))) ha" + show " (<(Q \* R) (s + s2)> \* <(s + s2 ## s')> \* I (s + s2 + s')) ha" + proof - + from hh have h0: "s ## s2 + s'" + by (metis pasrt_def sep_conjD) + with h1(2, 3) + have h2: "s + s2 ## s'" + by (metis sep_add_disjD sep_disj_addI1) + moreover from h1(2, 3) h2 have h3: "(s + (s2 + s')) = (s + s2 + s')" + by (metis `s ## s2 + s'` sep_add_assoc sep_add_disjD sep_disj_addD1) + moreover from hh have "Q s" by (metis pasrt_def sep_conjD) + moreover from h0 h1(2) h1(3) have "s ## s2" + by (metis sep_add_disjD sep_disj_addD) + moreover note h1(5) + ultimately show ?thesis + by (smt h0 hh sep_conjI) + qed + qed +qed + +lemma I_sequencing: + assumes h1: "I. \P\ c \Q\" + and h2: "I. \Q\ c \R\" + shows "I. \P\ c \R\" + using h1 h2 sequencing + by (smt IHoare_def) + +lemma I_pre_stren: + assumes h1: "I. \P\ c \Q\" + and h2: "\s. R s \ P s" + shows "I. \R\ c \Q\" +proof(unfold IHoare_def, default) + fix s' + show "\EXS s. \* <(s ## s')> \* I (s + s')\ c + \EXS s. \* <(s ## s')> \* I (s + s')\" + proof(rule pre_stren) + from h1[unfolded IHoare_def, rule_format, of s'] + show "\EXS s.

\* <(s ## s')> \* I (s + s')\ c + \EXS s. \* <(s ## s')> \* I (s + s')\" . + next + fix s + show "(EXS s. \* <(s ## s')> \* I (s + s')) s \ + (EXS s.

\* <(s ## s')> \* I (s + s')) s" + apply (unfold pred_ex_def, clarify) + apply (rule_tac x = sa in exI, sep_cancel+) + by (insert h2, auto simp:pasrt_def) + qed +qed + +lemma I_post_weaken: + assumes h1: "I. \P\ c \Q\" + and h2: "\ s. Q s \ R s" + shows "I. \P\ c \R\" +proof(unfold IHoare_def, default) + fix s' + show "\EXS s.

\* <(s ## s')> \* I (s + s')\ c + \EXS s. \* <(s ## s')> \* I (s + s')\" + proof(rule post_weaken) + from h1[unfolded IHoare_def, rule_format, of s'] + show "\EXS s.

\* <(s ## s')> \* I (s + s')\ c + \EXS s. \* <(s ## s')> \* I (s + s')\" . + next + fix s + show "(EXS s. \* <(s ## s')> \* I (s + s')) s \ + (EXS s. \* <(s ## s')> \* I (s + s')) s" + apply (unfold pred_ex_def, clarify) + apply (rule_tac x = sa in exI, sep_cancel+) + by (insert h2, auto simp:pasrt_def) + qed +qed + +lemma I_hoare_adjust: + assumes h1: "I. \P1\ c \Q1\" + and h2: "\s. P s \ P1 s" + and h3: "\s. Q1 s \ Q s" + shows "I. \P\ c \Q\" + using h1 h2 h3 I_post_weaken I_pre_stren + by (metis) + +lemma I_code_exI: + assumes h: "\ k. I. \P\ c(k) \Q\" + shows "I. \P\ EXS k. c(k) \Q\" +using h +by (smt IHoare_def code_exI) + +lemma I_code_extension: + assumes h: "I. \ P \ c \ Q \" + shows "I. \ P \ c ** e \ Q \" + using h + by (smt IHoare_def sep_exec.code_extension) + +lemma I_composition: + assumes h1: "I. \P\ c1 \Q\" + and h2: "I. \Q\ c2 \R\" + shows "I. \P\ c1 ** c2 \R\" + using h1 h2 +by (smt IHoare_def sep_exec.composition) + +lemma pre_condI: + assumes h: "cond \ \p\ c \q\" + shows "\ \* p\ c \q\" +proof(induct rule:HoareI) + case (Pre r s) + hence "cond" "(p \* c \* r) (recse s)" + apply (metis pasrt_def sep_conjD) + by (smt Pre.hyps pasrt_def sep_add_zero sep_conj_commute sep_conj_def) + from h[OF this(1), unfolded Hoare_gen_def, rule_format, OF this(2)] + show ?case . +qed + +lemma I_pre_condI: + assumes h: "cond \ I.\P\ c \Q\" + shows "I.\ \* P\ c \Q\" + using h +by (smt IHoareI condD cond_true_eq2 sep.mult_commute) + +lemma code_condI: + assumes h: "b \ \p\ c \q\" + shows "\p\ **c \q\" +proof(induct rule: HoareI) + case (Pre r s) + hence h1: "b" "(p \* c \* r) (recse s)" + apply (metis condD sep_conjD sep_conj_assoc) + by (smt Pre.hyps condD sep_conj_impl) + from h[OF h1(1), unfolded Hoare_gen_def, rule_format, OF h1(2)] + and h1(1) + show ?case + by (metis (full_types) cond_true_eq1) +qed + +lemma I_code_condI: + assumes h: "b \ I. \P\ c \Q\" + shows "I.\P\ **c \Q\" + using h +by (smt IHoareI condD cond_true_eq2 sep.mult_commute sep_conj_cond1) + +lemma precond_exI: + assumes h:"\x. \p x\ c \q\" + shows "\EXS x. p x\ c \q\" +proof(induct rule:HoareI) + case (Pre r s) + then obtain x where "(p x \* c \* r) (recse s)" + by (unfold pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI) + from h[of x, unfolded Hoare_gen_def, rule_format, OF this] + show ?case . +qed + +lemma I_precond_exI: + assumes h:"\x. I. \P x\ c \Q\" + shows "I.\EXS x. P x\ c \Q\" +proof(induct rule:IHoareI) + case (IPre s' s r cnf) + then obtain x + where "((EXS s.

\* <(s ## s')> \* I (s + s')) \* c \* r) (recse cnf)" + by ( auto elim!:sep_conjE intro!:sep_conjI simp:pred_ex_def pasrt_def sep_conj_ac) + from h[unfolded IHoare_def Hoare_gen_def, rule_format, OF this] + obtain k t + where "(( \* <(t ## s')> \* I (t + s')) \* c \* r) (recse (run (Suc k) cnf))" + by (unfold pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI) + thus ?case + by (auto simp:sep_conj_ac) +qed + +lemma hoare_sep_false: + "\sep_false\ c + \q\" + by(unfold Hoare_gen_def, clarify, simp) + +lemma I_hoare_sep_false: + "I. \sep_false\ c + \Q\" +by (smt IHoareI condD) + +end + +instantiation set :: (type)sep_algebra +begin + +definition set_zero_def: "0 = {}" + +definition plus_set_def: "s1 + s2 = s1 \ s2" + +definition sep_disj_set_def: "sep_disj s1 s2 = (s1 \ s2 = {})" + +lemmas set_ins_def = sep_disj_set_def plus_set_def set_zero_def + +instance + apply(default) + apply(simp add:set_ins_def) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + apply (metis inf_commute) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + apply (metis sup_commute) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + apply (metis (lifting) Un_assoc) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + apply (metis (lifting) Int_Un_distrib Un_empty inf_sup_distrib1 sup_eq_bot_iff) + apply(simp add:sep_disj_set_def plus_set_def set_zero_def) + by (metis (lifting) Int_Un_distrib Int_Un_distrib2 sup_eq_bot_iff) +end + +section {* A big operator of infinite separation conjunction *} + +definition "fam_conj I cpt s = (\ p. (s = (\ i \ I. p(i))) \ + (\ i \ I. cpt i (p i)) \ + (\ i \ I. \ j \ I. i \ j \ p(i) ## p(j)))" + +lemma fam_conj_zero_simp: "fam_conj {} cpt = " +proof + fix s + show "fam_conj {} cpt s = () s" + proof + assume "fam_conj {} cpt s" + then obtain p where "s = (\ i \ {}. p(i))" + by (unfold fam_conj_def, auto) + hence "s = {}" by auto + thus "() s" by (metis pasrt_def set_zero_def) + next + assume "() s" + hence eq_s: "s = {}" by (metis pasrt_def set_zero_def) + let ?p = "\ i. {}" + have "(s = (\ i \ {}. ?p(i)))" by (unfold eq_s, auto) + moreover have "(\ i \ {}. cpt i (?p i))" by auto + moreover have "(\ i \ {}. \ j \ {}. i \ j \ ?p(i) ## ?p(j))" by auto + ultimately show "fam_conj {} cpt s" + by (unfold eq_s fam_conj_def, auto) + qed +qed + +lemma fam_conj_disj_simp_pre: + assumes h1: "I = I1 + I2" + and h2: "I1 ## I2" + shows "fam_conj I cpt = (fam_conj I1 cpt \* fam_conj I2 cpt)" +proof + fix s + let ?fm = "fam_conj I cpt" and ?fm1 = "fam_conj I1 cpt" and ?fm2 = "fam_conj I2 cpt" + show "?fm s = (?fm1 \* ?fm2) s" + proof + assume "?fm s" + then obtain p where pre: + "s = (\ i \ I. p(i))" + "(\ i \ I. cpt i (p i))" + "(\ i \ I. \ j \ I. i \ j \ p(i) ## p(j))" + unfolding fam_conj_def by metis + from pre(1) h1 h2 have "s = (\ i \ I1. p(i)) + (\ i \ I2. p(i))" + by (auto simp:set_ins_def) + moreover from pre h1 have "?fm1 (\ i \ I1. p(i))" + by (unfold fam_conj_def, rule_tac x = p in exI, auto simp:set_ins_def) + moreover from pre h1 have "?fm2 (\ i \ I2. p(i))" + by (unfold fam_conj_def, rule_tac x = p in exI, auto simp:set_ins_def) + moreover have "(\ i \ I1. p(i)) ## (\ i \ I2. p(i))" + proof - + { fix x xa xb + assume h: "I1 \ I2 = {}" + "\i\I1 \ I2. \j\I1 \ I2. i \ j \ p i \ p j = {}" + "xa \ I1" "x \ p xa" "xb \ I2" "x \ p xb" + have "p xa \ p xb = {}" + proof(rule h(2)[rule_format]) + from h(1, 3, 5) show "xa \ I1 \ I2" by auto + next + from h(1, 3, 5) show "xb \ I1 \ I2" by auto + next + from h(1, 3, 5) show "xa \ xb" by auto + qed with h(4, 6) have False by auto + } with h1 h2 pre(3) show ?thesis by (auto simp:set_ins_def) + qed + ultimately show "(?fm1 \* ?fm2) s" by (auto intro!:sep_conjI) + next + assume "(?fm1 \* ?fm2) s" + then obtain s1 s2 where pre: + "s = s1 + s2" "s1 ## s2" "?fm1 s1" "?fm2 s2" + by (auto dest!:sep_conjD) + from pre(3) obtain p1 where pre1: + "s1 = (\ i \ I1. p1(i))" + "(\ i \ I1. cpt i (p1 i))" + "(\ i \ I1. \ j \ I1. i \ j \ p1(i) ## p1(j))" + unfolding fam_conj_def by metis + from pre(4) obtain p2 where pre2: + "s2 = (\ i \ I2. p2(i))" + "(\ i \ I2. cpt i (p2 i))" + "(\ i \ I2. \ j \ I2. i \ j \ p2(i) ## p2(j))" + unfolding fam_conj_def by metis + let ?p = "\ i. if i \ I1 then p1 i else p2 i" + from h2 pre(2) + have "s = (\ i \ I. ?p(i))" + apply (unfold h1 pre(1) pre1(1) pre2(1) set_ins_def, auto split:if_splits) + by (metis disjoint_iff_not_equal) + moreover from h2 pre1(2) pre2(2) have "(\ i \ I. cpt i (?p i))" + by (unfold h1 set_ins_def, auto split:if_splits) + moreover from pre1(1, 3) pre2(1, 3) pre(2) h2 + have "(\ i \ I. \ j \ I. i \ j \ ?p(i) ## ?p(j))" + apply (unfold h1 set_ins_def, auto split:if_splits) + by (metis IntI empty_iff) + ultimately show "?fm s" by (unfold fam_conj_def, auto) + qed +qed + +lemma fam_conj_disj_simp: + assumes h: "I1 ## I2" + shows "fam_conj (I1 + I2) cpt = (fam_conj I1 cpt \* fam_conj I2 cpt)" +proof - + from fam_conj_disj_simp_pre[OF _ h, of "I1 + I2"] + show ?thesis by simp +qed + +lemma fam_conj_elm_simp: + assumes h: "i \ I" + shows "fam_conj I cpt = (cpt(i) \* fam_conj (I - {i}) cpt)" +proof + fix s + show "fam_conj I cpt s = (cpt i \* fam_conj (I - {i}) cpt) s" + proof + assume pre: "fam_conj I cpt s" + show "(cpt i \* fam_conj (I - {i}) cpt) s" + proof - + from pre obtain p where pres: + "s = (\ i \ I. p(i))" + "(\ i \ I. cpt i (p i))" + "(\ i \ I. \ j \ I. i \ j \ p(i) ## p(j))" + unfolding fam_conj_def by metis + let ?s = "(\i\(I - {i}). p i)" + from pres(3) h have disj: "(p i) ## ?s" + by (auto simp:set_ins_def) + moreover from pres(1) h have eq_s: "s = (p i) + ?s" + by (auto simp:set_ins_def) + moreover from pres(2) h have "cpt i (p i)" by auto + moreover from pres have "(fam_conj (I - {i}) cpt) ?s" + by (unfold fam_conj_def, rule_tac x = p in exI, auto) + ultimately show ?thesis by (auto intro!:sep_conjI) + qed + next + let ?fam = "fam_conj (I - {i}) cpt" + assume "(cpt i \* ?fam) s" + then obtain s1 s2 where pre: + "s = s1 + s2" "s1 ## s2" "cpt i s1" "?fam s2" + by (auto dest!:sep_conjD) + from pre(4) obtain p where pres: + "s2 = (\ ia \ I - {i}. p(ia))" + "(\ ia \ I - {i}. cpt ia (p ia))" + "(\ ia \ I - {i}. \ j \ I - {i}. ia \ j \ p(ia) ## p(j))" + unfolding fam_conj_def by metis + let ?p = "p(i:=s1)" + from h pres(3) pre(2)[unfolded pres(1)] + have " \ia\I. \j\I. ia \ j \ ?p ia ## ?p j" by (auto simp:set_ins_def) + moreover from pres(1) pre(1) h pre(2) + have "(s = (\ i \ I. ?p(i)))" by (auto simp:set_ins_def split:if_splits) + moreover from pre(3) pres(2) h + have "(\ i \ I. cpt i (?p i))" by (auto simp:set_ins_def split:if_splits) + ultimately show "fam_conj I cpt s" + by (unfold fam_conj_def, auto) + qed +qed + +lemma fam_conj_insert_simp: + assumes h:"i \ I" + shows "fam_conj (insert i I) cpt = (cpt(i) \* fam_conj I cpt)" +proof - + have "i \ insert i I" by auto + from fam_conj_elm_simp[OF this] and h + show ?thesis by auto +qed + +lemmas fam_conj_simps = fam_conj_insert_simp fam_conj_zero_simp + +lemma "fam_conj {0, 2, 3} cpt = (cpt 2 \* cpt (0::int) \* cpt 3)" + by (simp add:fam_conj_simps sep_conj_ac) + +lemma fam_conj_ext_eq: + assumes h: "\ i . i \ I \ f i = g i" + shows "fam_conj I f = fam_conj I g" +proof + fix s + show "fam_conj I f s = fam_conj I g s" + by (unfold fam_conj_def, auto simp:h) +qed + +end +