# HG changeset patch # User wenzelm # Date 1148923384 -7200 # Node ID 163f1ba9225a69b62fd1e3e9ded47c88542077ce # Parent 9ac97dc142148c4dbc6c312af5b0adf2f3193f03 tuned; diff -r 9ac97dc14214 -r 163f1ba9225a src/HOL/W0/W0.thy --- a/src/HOL/W0/W0.thy Mon May 29 17:38:30 2006 +0200 +++ b/src/HOL/W0/W0.thy Mon May 29 19:23:04 2006 +0200 @@ -425,8 +425,7 @@ section {* Correctness and completeness of the type inference algorithm W *} consts - W :: "expr \ typ list \ nat \ (subst \ typ \ nat) maybe" ("\") - + "\" :: "expr \ typ list \ nat \ (subst \ typ \ nat) maybe" primrec "\ (Var i) a n = (if i < length a then Ok (id_subst, a ! i, n) else Fail)" @@ -468,14 +467,14 @@ by (simp add: subst_comp_tel o_def) show "$s a |- e1 :: $u t2 -> t" proof - - from W1_ok have "$s1 a |- e1 :: t1" by (rule App.hyps) + from W1_ok have "$s1 a |- e1 :: t1" by (rule App.hyps(1)) then have "$u ($s2 ($s1 a)) |- e1 :: $u ($s2 t1)" by (intro has_type_subst_closed) with s' t mgu_ok show ?thesis by simp qed show "$s a |- e2 :: $u t2" proof - - from W2_ok have "$s2 ($s1 a) |- e2 :: t2" by (rule App.hyps) + from W2_ok have "$s2 ($s1 a) |- e2 :: t2" by (rule App.hyps(2)) then have "$u ($s2 ($s1 a)) |- e2 :: $u t2" by (rule has_type_subst_closed) with s' show ?thesis by simp @@ -785,7 +784,7 @@ *} consts - I :: "expr \ typ list \ nat \ subst \ (subst \ typ \ nat) maybe" ("\") + "\" :: "expr \ typ list \ nat \ subst \ (subst \ typ \ nat) maybe" primrec "\ (Var i) a n s = (if i < length a then Ok (s, a ! i, n) else Fail)" "\ (Abs e) a n s = ((s, t, m) := \ e (TVar n # a) (Suc n) s; diff -r 9ac97dc14214 -r 163f1ba9225a src/ZF/IMP/Denotation.thy --- a/src/ZF/IMP/Denotation.thy Mon May 29 17:38:30 2006 +0200 +++ b/src/ZF/IMP/Denotation.thy Mon May 29 19:23:04 2006 +0200 @@ -14,9 +14,9 @@ B :: "i => i => i" C :: "i => i" -constdefs +definition Gamma :: "[i,i,i] => i" ("\") - "\(b,cden) == + "\(b,cden) = (\phi. {io \ (phi O cden). B(b,fst(io))=1} \ {io \ id(loc->nat). B(b,fst(io))=0})"