diff -r 0b21101157b1 -r 6853ce305118 Nominal/TySch.thy --- a/Nominal/TySch.thy Fri Mar 19 09:03:10 2010 +0100 +++ b/Nominal/TySch.thy Fri Mar 19 09:31:38 2010 +0100 @@ -61,17 +61,22 @@ lemma induct: assumes a1: "\name b. P b (Var name)" and a2: "\t1 t2 b. \\c. P c t1; \c. P c t2\ \ P b (Fun t1 t2)" - and a3: "\fset t. \\c. P c t; fset_to_set (fmap atom fset) \* b\ \ P' b (All fset t)" - shows "P a t" + and a3: "\fset t b. \\c. P c t; fset_to_set (fmap atom fset) \* b\ \ P' b (All fset t)" + shows "P a t \ P' d ts " proof - - have "\p. P a (p \ t)" - apply (induct t rule: t_tyS.inducts(1)) + have " (\p. P a (p \ t)) \ (\p. P' d (p \ ts))" + apply (rule t_tyS.induct) apply (simp add: a1) apply (simp_all) - apply (rule_tac ?t1.0="p \ t1" and ?t2.0="p \ t2" in a2) + apply (rule allI) + apply (rule a2) + defer defer + apply (rule allI) + apply (rule a3) + apply simp_all sorry - then have "P a (0 \ t)" by blast - then show "P a t" by simp + then have "P a (0 \ t) \ P' d (0 \ ts)" by blast + then show ?thesis by simp qed lemma