diff -r 11bcd77cfa22 -r b1f10b98430d src/HOL/Induct/LList.thy --- a/src/HOL/Induct/LList.thy Wed Oct 12 03:02:18 2005 +0200 +++ b/src/HOL/Induct/LList.thy Wed Oct 12 10:49:07 2005 +0200 @@ -143,8 +143,8 @@ declare option.split [split] text{*This justifies using llist in other recursive type definitions*} -lemma llist_mono: "A<=B ==> llist(A) <= llist(B)" -apply (unfold llist.defs ) +lemma llist_mono: "A\B ==> llist(A) \ llist(B)" +apply (simp add: llist.defs) apply (rule gfp_mono) apply (assumption | rule basic_monos)+ done @@ -163,40 +163,36 @@ *} lemma llist_coinduct: - "[| M \ X; X <= list_Fun A (X Un llist(A)) |] ==> M \ llist(A)" -apply (unfold list_Fun_def) + "[| M \ X; X \ list_Fun A (X Un llist(A)) |] ==> M \ llist(A)" +apply (simp add: list_Fun_def) apply (erule llist.coinduct) -apply (erule subsetD [THEN CollectD], assumption) +apply (blast intro: elim:); done lemma list_Fun_NIL_I [iff]: "NIL \ list_Fun A X" -by (unfold list_Fun_def NIL_def, fast) +by (simp add: list_Fun_def NIL_def) lemma list_Fun_CONS_I [intro!,simp]: "[| M \ A; N \ X |] ==> CONS M N \ list_Fun A X" -apply (unfold list_Fun_def CONS_def, fast) -done +by (simp add: list_Fun_def CONS_def) text{*Utilise the ``strong'' part, i.e. @{text "gfp(f)"}*} lemma list_Fun_llist_I: "M \ llist(A) ==> M \ list_Fun A (X Un llist(A))" apply (unfold llist.defs list_Fun_def) apply (rule gfp_fun_UnI2) -apply (rule monoI, blast) -apply assumption +apply (rule monoI, auto) done subsection{* @{text LList_corec} satisfies the desired recurion equation *} text{*A continuity result?*} lemma CONS_UN1: "CONS M (\x. f(x)) = (\x. CONS M (f x))" -apply (unfold CONS_def) -apply (simp add: In1_UN1 Scons_UN1_y) +apply (simp add: CONS_def In1_UN1 Scons_UN1_y) done -lemma CONS_mono: "[| M<=M'; N<=N' |] ==> CONS M N <= CONS M' N'" -apply (unfold CONS_def) -apply (assumption | rule In1_mono Scons_mono)+ +lemma CONS_mono: "[| M\M'; N\N' |] ==> CONS M N \ CONS M' N'" +apply (simp add: CONS_def In1_mono Scons_mono) done declare LList_corec_fun_def [THEN def_nat_rec_0, simp] @@ -205,19 +201,19 @@ subsubsection{* The directions of the equality are proved separately *} lemma LList_corec_subset1: - "LList_corec a f <= + "LList_corec a f \ (case f a of None => NIL | Some(z,w) => CONS z (LList_corec w f))" apply (unfold LList_corec_def) apply (rule UN_least) -apply (case_tac "k") +apply (case_tac k) apply (simp_all (no_asm_simp)) apply (rule allI impI subset_refl [THEN CONS_mono] UNIV_I [THEN UN_upper])+ done lemma LList_corec_subset2: - "(case f a of None => NIL | Some(z,w) => CONS z (LList_corec w f)) <= + "(case f a of None => NIL | Some(z,w) => CONS z (LList_corec w f)) \ LList_corec a f" -apply (unfold LList_corec_def) +apply (simp add: LList_corec_def) apply (simp add: CONS_UN1, safe) apply (rule_tac a="Suc(?k)" in UN_I, simp, simp)+ done @@ -237,7 +233,7 @@ text{*A typical use of co-induction to show membership in the @{text gfp}. Bisimulation is @{text "range(%x. LList_corec x f)"} *} lemma LList_corec_type: "LList_corec a f \ llist UNIV" -apply (rule_tac X = "range (%x. LList_corec x ?g) " in llist_coinduct) +apply (rule_tac X = "range (%x. LList_corec x ?g)" in llist_coinduct) apply (rule rangeI, safe) apply (subst LList_corec, simp) done @@ -252,7 +248,7 @@ lemma LListD_implies_ntrunc_equality [rule_format]: "\M N. (M,N) \ LListD(diag A) --> ntrunc k M = ntrunc k N" -apply (induct_tac "k" rule: nat_less_induct) +apply (induct_tac "k" rule: nat_less_induct) apply (safe del: equalityI) apply (erule LListD.cases) apply (safe del: equalityI) @@ -264,20 +260,19 @@ text{*The domain of the @{text LListD} relation*} lemma Domain_LListD: - "Domain (LListD(diag A)) <= llist(A)" -apply (unfold llist.defs NIL_def CONS_def) + "Domain (LListD(diag A)) \ llist(A)" +apply (simp add: llist.defs NIL_def CONS_def) apply (rule gfp_upperbound) txt{*avoids unfolding @{text LListD} on the rhs*} -apply (rule_tac P = "%x. Domain x <= ?B" in LListD_unfold [THEN ssubst], simp) -apply blast +apply (rule_tac P = "%x. Domain x \ ?B" in LListD_unfold [THEN ssubst], auto) done text{*This inclusion justifies the use of coinduction to show @{text "M = N"}*} -lemma LListD_subset_diag: "LListD(diag A) <= diag(llist(A))" +lemma LListD_subset_diag: "LListD(diag A) \ diag(llist(A))" apply (rule subsetI) -apply (rule_tac p = "x" in PairE, safe) +apply (rule_tac p = x in PairE, safe) apply (rule diag_eqI) -apply (rule LListD_implies_ntrunc_equality [THEN ntrunc_equality], assumption) +apply (rule LListD_implies_ntrunc_equality [THEN ntrunc_equality], assumption) apply (erule DomainI [THEN Domain_LListD [THEN subsetD]]) done @@ -286,38 +281,36 @@ text {* THE COINDUCTIVE DEFINITION PACKAGE COULD DO THIS! *} -lemma LListD_Fun_mono: "A<=B ==> LListD_Fun r A <= LListD_Fun r B" -apply (unfold LListD_Fun_def) +lemma LListD_Fun_mono: "A\B ==> LListD_Fun r A \ LListD_Fun r B" +apply (simp add: LListD_Fun_def) apply (assumption | rule basic_monos)+ done lemma LListD_coinduct: - "[| M \ X; X <= LListD_Fun r (X Un LListD(r)) |] ==> M \ LListD(r)" -apply (unfold LListD_Fun_def) + "[| M \ X; X \ LListD_Fun r (X Un LListD(r)) |] ==> M \ LListD(r)" +apply (simp add: LListD_Fun_def) apply (erule LListD.coinduct) -apply (erule subsetD [THEN CollectD], assumption) +apply (auto ); done lemma LListD_Fun_NIL_I: "(NIL,NIL) \ LListD_Fun r s" -by (unfold LListD_Fun_def NIL_def, fast) +by (simp add: LListD_Fun_def NIL_def) lemma LListD_Fun_CONS_I: "[| x\A; (M,N):s |] ==> (CONS x M, CONS x N) \ LListD_Fun (diag A) s" -apply (unfold LListD_Fun_def CONS_def, blast) -done +by (simp add: LListD_Fun_def CONS_def, blast) text{*Utilise the "strong" part, i.e. @{text "gfp(f)"}*} lemma LListD_Fun_LListD_I: "M \ LListD(r) ==> M \ LListD_Fun r (X Un LListD(r))" apply (unfold LListD.defs LListD_Fun_def) apply (rule gfp_fun_UnI2) -apply (rule monoI, blast) -apply assumption +apply (rule monoI, auto) done text{*This converse inclusion helps to strengthen @{text LList_equalityI}*} -lemma diag_subset_LListD: "diag(llist(A)) <= LListD(diag A)" +lemma diag_subset_LListD: "diag(llist(A)) \ LListD(diag A)" apply (rule subsetI) apply (erule LListD_coinduct) apply (rule subsetI) @@ -342,7 +335,7 @@ [also admits true equality] Replace @{text A} by some particular set, like @{text "{x. True}"}??? *} lemma LList_equalityI: - "[| (M,N) \ r; r <= LListD_Fun (diag A) (r Un diag(llist(A))) |] + "[| (M,N) \ r; r \ LListD_Fun (diag A) (r Un diag(llist(A))) |] ==> M=N" apply (rule LListD_subset_diag [THEN subsetD, THEN diagE]) apply (erule LListD_coinduct) @@ -365,7 +358,7 @@ ==> h1=h2" apply (rule ext) txt{*next step avoids an unknown (and flexflex pair) in simplification*} -apply (rule_tac A = "UNIV" and r = "range(%u. (h1 u, h2 u))" +apply (rule_tac A = UNIV and r = "range(%u. (h1 u, h2 u))" in LList_equalityI) apply (rule rangeI, safe) apply (simp add: LListD_Fun_NIL_I UNIV_I [THEN LListD_Fun_CONS_I]) @@ -395,7 +388,7 @@ "!!x. h2 x = (case f x of None => NIL | Some(z,w) => CONS z (h2 w))" shows "h1=h2" apply (rule ntrunc_equality [THEN ext]) -apply (rule_tac x = "x" in spec) +apply (rule_tac x = x in spec) apply (induct_tac "k" rule: nat_less_induct) apply (rename_tac "n") apply (rule allI) @@ -404,8 +397,7 @@ apply (intro strip) apply (case_tac "n") apply (rename_tac [2] "m") -apply (case_tac [2] "m") -apply simp_all +apply (case_tac [2] "m", simp_all) done @@ -444,16 +436,16 @@ subsection{* Isomorphisms *} lemma LListI: "x \ llist (range Leaf) ==> x \ LList" -by (unfold LList_def, simp) +by (simp add: LList_def) lemma LListD: "x \ LList ==> x \ llist (range Leaf)" -by (unfold LList_def, simp) +by (simp add: LList_def) subsubsection{* Distinctness of constructors *} lemma LCons_not_LNil [iff]: "~ LCons x xs = LNil" -apply (unfold LNil_def LCons_def) +apply (simp add: LNil_def LCons_def) apply (subst Abs_LList_inject) apply (rule llist.intros CONS_not_NIL rangeI LListI Rep_LList [THEN LListD])+ done @@ -464,12 +456,12 @@ subsubsection{* llist constructors *} lemma Rep_LList_LNil: "Rep_LList LNil = NIL" -apply (unfold LNil_def) +apply (simp add: LNil_def) apply (rule llist.NIL_I [THEN LListI, THEN Abs_LList_inverse]) done lemma Rep_LList_LCons: "Rep_LList(LCons x l) = CONS (Leaf x) (Rep_LList l)" -apply (unfold LCons_def) +apply (simp add: LCons_def) apply (rule llist.CONS_I [THEN LListI, THEN Abs_LList_inverse] rangeI Rep_LList [THEN LListD])+ done @@ -477,8 +469,7 @@ subsubsection{* Injectiveness of @{text CONS} and @{text LCons} *} lemma CONS_CONS_eq2: "(CONS M N=CONS M' N') = (M=M' & N=N')" -apply (unfold CONS_def) -apply (fast elim!: Scons_inject) +apply (simp add: CONS_def) done lemmas CONS_inject = CONS_CONS_eq [THEN iffD1, THEN conjE, standard] @@ -490,7 +481,7 @@ declare llist.intros [intro] lemma LCons_LCons_eq [iff]: "(LCons x xs=LCons y ys) = (x=y & xs=ys)" -apply (unfold LCons_def) +apply (simp add: LCons_def) apply (subst Abs_LList_inject) apply (auto simp add: Rep_LList_inject) done @@ -546,7 +537,7 @@ subsubsection{* Two easy results about @{text Lmap} *} lemma Lmap_compose: "M \ llist(A) ==> Lmap (f o g) M = Lmap f (Lmap g M)" -apply (unfold o_def) +apply (simp add: o_def) apply (drule imageI) apply (erule LList_equalityI, safe) apply (erule llist.cases, simp_all) @@ -566,19 +557,19 @@ subsection{* @{text Lappend} -- its two arguments cause some complications! *} lemma Lappend_NIL_NIL [simp]: "Lappend NIL NIL = NIL" -apply (unfold Lappend_def) +apply (simp add: Lappend_def) apply (rule LList_corec [THEN trans], simp) done lemma Lappend_NIL_CONS [simp]: "Lappend NIL (CONS N N') = CONS N (Lappend NIL N')" -apply (unfold Lappend_def) +apply (simp add: Lappend_def) apply (rule LList_corec [THEN trans], simp) done lemma Lappend_CONS [simp]: "Lappend (CONS M M') N = CONS M (Lappend M' N)" -apply (unfold Lappend_def) +apply (simp add: Lappend_def) apply (rule LList_corec [THEN trans], simp) done @@ -600,17 +591,16 @@ apply (rule_tac X = "\u\llist (A) . \v \ llist (A) . {Lappend u v}" in llist_coinduct) apply fast apply safe -apply (erule_tac aa = "u" in llist.cases) -apply (erule_tac aa = "v" in llist.cases, simp_all) -apply blast +apply (erule_tac aa = u in llist.cases) +apply (erule_tac aa = v in llist.cases, simp_all, blast) done text{*strong co-induction: bisimulation and case analysis on one variable*} lemma Lappend_type': "[| M \ llist(A); N \ llist(A) |] ==> Lappend M N \ llist(A)" -apply (rule_tac X = " (%u. Lappend u N) `llist (A) " in llist_coinduct) +apply (rule_tac X = "(%u. Lappend u N) `llist (A)" in llist_coinduct) apply (erule imageI) apply (rule image_subsetI) -apply (erule_tac aa = "x" in llist.cases) +apply (erule_tac aa = x in llist.cases) apply (simp add: list_Fun_llist_I, simp) done @@ -624,18 +614,16 @@ declare rangeI [simp] inj_Leaf [simp] lemma llist_case_LNil [simp]: "llist_case c d LNil = c" -by (unfold llist_case_def LNil_def, simp) +by (simp add: llist_case_def LNil_def) -lemma llist_case_LCons [simp]: - "llist_case c d (LCons M N) = d M N" -apply (unfold llist_case_def LCons_def, simp) -done +lemma llist_case_LCons [simp]: "llist_case c d (LCons M N) = d M N" +by (simp add: llist_case_def LCons_def) + text{*Elimination is case analysis, not induction.*} lemma llistE: "[| l=LNil ==> P; !!x l'. l=LCons x l' ==> P |] ==> P" apply (rule Rep_LList [THEN LListD, THEN llist.cases]) - apply (simp add: Rep_LList_LNil [symmetric] Rep_LList_inject) - apply blast + apply (simp add: Rep_LList_LNil [symmetric] Rep_LList_inject, blast) apply (erule LListI [THEN Rep_LList_cases], clarify) apply (simp add: Rep_LList_LCons [symmetric] Rep_LList_inject, blast) done @@ -647,7 +635,7 @@ "LList_corec a (%z. case f z of None => None | Some(v,w) => Some(Leaf(v),w)) \ llist(range Leaf)" -apply (rule_tac X = "range (%x. LList_corec x ?g) " in llist_coinduct) +apply (rule_tac X = "range (%x. LList_corec x ?g)" in llist_coinduct) apply (rule rangeI, safe) apply (subst LList_corec, force) done @@ -673,18 +661,18 @@ subsection{* Deriving @{text llist_equalityI} -- @{text llist} equality is a bisimulation *} lemma LListD_Fun_subset_Times_llist: - "r <= (llist A) <*> (llist A) - ==> LListD_Fun (diag A) r <= (llist A) <*> (llist A)" + "r \ (llist A) <*> (llist A) + ==> LListD_Fun (diag A) r \ (llist A) <*> (llist A)" by (auto simp add: LListD_Fun_def) lemma subset_Times_llist: - "prod_fun Rep_LList Rep_LList ` r <= + "prod_fun Rep_LList Rep_LList ` r \ (llist(range Leaf)) <*> (llist(range Leaf))" by (blast intro: Rep_LList [THEN LListD]) lemma prod_fun_lemma: - "r <= (llist(range Leaf)) <*> (llist(range Leaf)) - ==> prod_fun (Rep_LList o Abs_LList) (Rep_LList o Abs_LList) ` r <= r" + "r \ (llist(range Leaf)) <*> (llist(range Leaf)) + ==> prod_fun (Rep_LList o Abs_LList) (Rep_LList o Abs_LList) ` r \ r" apply safe apply (erule subsetD [THEN SigmaE2], assumption) apply (simp add: LListI [THEN Abs_LList_inverse]) @@ -699,8 +687,8 @@ text{*Used with @{text lfilter}*} lemma llistD_Fun_mono: - "A<=B ==> llistD_Fun A <= llistD_Fun B" -apply (unfold llistD_Fun_def prod_fun_def, auto) + "A\B ==> llistD_Fun A \ llistD_Fun B" +apply (simp add: llistD_Fun_def prod_fun_def, auto) apply (rule image_eqI) prefer 2 apply (blast intro: rev_subsetD [OF _ LListD_Fun_mono], force) done @@ -708,10 +696,10 @@ subsubsection{* To show two llists are equal, exhibit a bisimulation! [also admits true equality] *} lemma llist_equalityI: - "[| (l1,l2) \ r; r <= llistD_Fun(r Un range(%x.(x,x))) |] ==> l1=l2" -apply (unfold llistD_Fun_def) + "[| (l1,l2) \ r; r \ llistD_Fun(r Un range(%x.(x,x))) |] ==> l1=l2" +apply (simp add: llistD_Fun_def) apply (rule Rep_LList_inject [THEN iffD1]) -apply (rule_tac r = "prod_fun Rep_LList Rep_LList `r" and A = "range (Leaf) " in LList_equalityI) +apply (rule_tac r = "prod_fun Rep_LList Rep_LList `r" and A = "range (Leaf)" in LList_equalityI) apply (erule prod_fun_imageI) apply (erule image_mono [THEN subset_trans]) apply (rule image_compose [THEN subst]) @@ -725,20 +713,20 @@ subsubsection{* Rules to prove the 2nd premise of @{text llist_equalityI} *} lemma llistD_Fun_LNil_I [simp]: "(LNil,LNil) \ llistD_Fun(r)" -apply (unfold llistD_Fun_def LNil_def) +apply (simp add: llistD_Fun_def LNil_def) apply (rule LListD_Fun_NIL_I [THEN prod_fun_imageI]) done lemma llistD_Fun_LCons_I [simp]: "(l1,l2):r ==> (LCons x l1, LCons x l2) \ llistD_Fun(r)" -apply (unfold llistD_Fun_def LCons_def) +apply (simp add: llistD_Fun_def LCons_def) apply (rule rangeI [THEN LListD_Fun_CONS_I, THEN prod_fun_imageI]) apply (erule prod_fun_imageI) done text{*Utilise the "strong" part, i.e. @{text "gfp(f)"}*} lemma llistD_Fun_range_I: "(l,l) \ llistD_Fun(r Un range(%x.(x,x)))" -apply (unfold llistD_Fun_def) +apply (simp add: llistD_Fun_def) apply (rule Rep_LList_inverse [THEN subst]) apply (rule prod_fun_imageI) apply (subst image_Un) @@ -752,9 +740,9 @@ !!x l. (f(LCons x l),g(LCons x l)) \ llistD_Fun(range(%u. (f(u),g(u))) Un range(%v. (v,v))) |] ==> f(l) = (g(l :: 'a llist) :: 'b llist)" -apply (rule_tac r = "range (%u. (f (u),g (u))) " in llist_equalityI) +apply (rule_tac r = "range (%u. (f (u),g (u)))" in llist_equalityI) apply (rule rangeI, clarify) -apply (rule_tac l = "u" in llistE) +apply (rule_tac l = u in llistE) apply (simp_all add: llistD_Fun_range_I) done @@ -771,10 +759,10 @@ subsubsection{* Two easy results about @{text lmap} *} lemma lmap_compose [simp]: "lmap (f o g) l = lmap f (lmap g l)" -by (rule_tac l = "l" in llist_fun_equalityI, simp_all) +by (rule_tac l = l in llist_fun_equalityI, simp_all) lemma lmap_ident [simp]: "lmap (%x. x) l = l" -by (rule_tac l = "l" in llist_fun_equalityI, simp_all) +by (rule_tac l = l in llist_fun_equalityI, simp_all) subsection{* iterates -- @{text llist_fun_equalityI} cannot be used! *} @@ -783,10 +771,10 @@ by (rule iterates_def [THEN def_llist_corec, THEN trans], simp) lemma lmap_iterates [simp]: "lmap f (iterates f x) = iterates f (f x)" -apply (rule_tac r = "range (%u. (lmap f (iterates f u),iterates f (f u))) " in llist_equalityI) +apply (rule_tac r = "range (%u. (lmap f (iterates f u),iterates f (f u)))" in llist_equalityI) apply (rule rangeI, safe) -apply (rule_tac x1 = "f (u) " in iterates [THEN ssubst]) -apply (rule_tac x1 = "u" in iterates [THEN ssubst], simp) +apply (rule_tac x1 = "f (u)" in iterates [THEN ssubst]) +apply (rule_tac x1 = u in iterates [THEN ssubst], simp) done lemma iterates_lmap: "iterates f x = LCons x (lmap f (iterates f x))" @@ -801,8 +789,7 @@ lemma fun_power_lmap: "nat_rec (LCons b l) (%m. lmap(f)) n = LCons (nat_rec b (%m. f) n) (nat_rec l (%m. lmap(f)) n)" -apply (induct_tac "n", simp_all) -done +by (induct_tac "n", simp_all) lemma fun_power_Suc: "nat_rec (g x) (%m. g) n = nat_rec x (%m. g) (Suc n)" by (induct_tac "n", simp_all) @@ -836,35 +823,34 @@ subsection{* @{text lappend} -- its two arguments cause some complications! *} lemma lappend_LNil_LNil [simp]: "lappend LNil LNil = LNil" -apply (unfold lappend_def) +apply (simp add: lappend_def) apply (rule llist_corec [THEN trans], simp) done lemma lappend_LNil_LCons [simp]: "lappend LNil (LCons l l') = LCons l (lappend LNil l')" -apply (unfold lappend_def) +apply (simp add: lappend_def) apply (rule llist_corec [THEN trans], simp) done lemma lappend_LCons [simp]: "lappend (LCons l l') N = LCons l (lappend l' N)" -apply (unfold lappend_def) +apply (simp add: lappend_def) apply (rule llist_corec [THEN trans], simp) done lemma lappend_LNil [simp]: "lappend LNil l = l" -by (rule_tac l = "l" in llist_fun_equalityI, simp_all) +by (rule_tac l = l in llist_fun_equalityI, simp_all) lemma lappend_LNil2 [simp]: "lappend l LNil = l" -by (rule_tac l = "l" in llist_fun_equalityI, simp_all) +by (rule_tac l = l in llist_fun_equalityI, simp_all) text{*The infinite first argument blocks the second*} lemma lappend_iterates [simp]: "lappend (iterates f x) N = iterates f x" apply (rule_tac r = "range (%u. (lappend (iterates f u) N,iterates f u))" in llist_equalityI) - apply (rule rangeI) -apply (safe) + apply (rule rangeI, safe) apply (subst (1 2) iterates) apply simp done @@ -879,22 +865,18 @@ in llist_equalityI) apply (rule UN1_I) apply (rule rangeI, safe) -apply (rule_tac l = "l" in llistE) -apply (rule_tac l = "n" in llistE, simp_all) +apply (rule_tac l = l in llistE) +apply (rule_tac l = n in llistE, simp_all) apply (blast intro: llistD_Fun_LCons_I) done text{*Shorter proof of theorem above using @{text llist_equalityI} as strong coinduction*} lemma lmap_lappend_distrib': "lmap f (lappend l n) = lappend (lmap f l) (lmap f n)" -apply (rule_tac l = "l" in llist_fun_equalityI, simp) -apply simp -done +by (rule_tac l = l in llist_fun_equalityI, auto) text{*Without strong coinduction, three case analyses might be needed*} lemma lappend_assoc': "lappend (lappend l1 l2) l3 = lappend l1 (lappend l2 l3)" -apply (rule_tac l = "l1" in llist_fun_equalityI, simp) -apply simp -done +by (rule_tac l = l1 in llist_fun_equalityI, auto) end