# HG changeset patch # User nipkow # Date 814610806 -3600 # Node ID c7a8f374339b1bd4c12718ae686d26b75836865c # Parent e74f694ca1da01b0e5cf507fba7763aee4f1377d New theory: type inference for let-free MiniML diff -r e74f694ca1da -r c7a8f374339b src/HOL/MiniML/I.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MiniML/I.ML Wed Oct 25 09:46:46 1995 +0100 @@ -0,0 +1,123 @@ +open I; + +Unify.trace_bound := 45; +Unify.search_bound := 50; + +goal thy + "! a m s s' t n. \ +\ (new_tv m a & new_tv m s) --> I e a m s = Ok(s',t,n) --> \ +\ ( ? r. W e ($ s a) m = Ok(r, $ s' t, n) & s' = ($ r o s) )"; +by (expr.induct_tac "e" 1); + +(* case Var n *) +by (simp_tac (!simpset addsimps [app_subst_list] + setloop (split_tac [expand_if])) 1); +by (fast_tac (HOL_cs addss !simpset) 1); + +(* case Abs e *) +by (asm_full_simp_tac (!simpset setloop (split_tac [expand_bind])) 1); +by (strip_tac 1); +br conjI 1; + +by (strip_tac 1); +by (REPEAT (etac allE 1)); +be impE 1; +by (fast_tac (HOL_cs addss (!simpset addsimps [new_tv_subst])) 2); +by (fast_tac (HOL_cs addIs [new_tv_Suc_list RS mp,new_tv_subst_le, + less_imp_le,lessI]) 1); + +by (strip_tac 1); +by (REPEAT (etac allE 1)); +be impE 1; +by (fast_tac (HOL_cs addss (!simpset addsimps [new_tv_subst])) 2); +by (fast_tac (HOL_cs addIs [new_tv_Suc_list RS mp,new_tv_subst_le, + less_imp_le,lessI]) 1); + +(* case App e1 e2 *) +by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1); +by (strip_tac 1); +by (rename_tac "s1' t1 n1 s2' t2 n2 sa" 1); +br conjI 1; +by (fast_tac (HOL_cs addss !simpset) 1); + +by (strip_tac 1); +by (rename_tac "s1 t1' n1'" 1); +by (eres_inst_tac [("x","a")] allE 1); +by (eres_inst_tac [("x","m")] allE 1); +by (eres_inst_tac [("x","s")] allE 1); +by (eres_inst_tac [("x","s1'")] allE 1); +by (eres_inst_tac [("x","t1")] allE 1); +by (eres_inst_tac [("x","n1")] allE 1); +by (eres_inst_tac [("x","a")] allE 1); +by (eres_inst_tac [("x","n1")] allE 1); +by (eres_inst_tac [("x","s1'")] allE 1); +by (eres_inst_tac [("x","s2'")] allE 1); +by (eres_inst_tac [("x","t2")] allE 1); +by (eres_inst_tac [("x","n2")] allE 1); + +br conjI 1; +by (strip_tac 1); +by (mp_tac 1); +by (mp_tac 1); +be exE 1; +be conjE 1; +be impE 1; +by ((forward_tac [new_tv_subst_tel] 1) THEN (atac 1)); +by ((dres_inst_tac [("xc","$ s a")] new_tv_W 1) THEN (atac 1)); +by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] + addss !simpset) 1); +by (fast_tac (HOL_cs addss (!simpset addsimps [subst_comp_tel])) 1); + +by (strip_tac 1); +by (rename_tac "s2 t2' n2'" 1); +br conjI 1; +by (strip_tac 1); +by (mp_tac 1); +by (mp_tac 1); +be exE 1; +be conjE 1; +be impE 1; +by ((forward_tac [new_tv_subst_tel] 1) THEN (atac 1)); +by ((dres_inst_tac [("xc","$ s a")] new_tv_W 1) THEN (atac 1)); +by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] + addss !simpset) 1); +by (fast_tac (HOL_cs addss (!simpset addsimps [subst_comp_tel,subst_comp_te])) 1); + +by (strip_tac 1); +by (mp_tac 1); +by (mp_tac 1); +be exE 1; +be conjE 1; +be impE 1; +by ((forward_tac [new_tv_subst_tel] 1) THEN (atac 1)); +by ((dres_inst_tac [("xc","$ s a")] new_tv_W 1) THEN (atac 1)); +by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] + addss !simpset) 1); + +by (mp_tac 1); +by (REPEAT (eresolve_tac [exE,conjE] 1)); +by (REPEAT + ((asm_full_simp_tac (!simpset addsimps [subst_comp_tel,subst_comp_te]) 1) THEN + (REPEAT (etac conjE 1)) THEN (hyp_subst_tac 1) )); +br exI 1; +by (safe_tac HOL_cs); +by (fast_tac HOL_cs 1); +by (Simp_tac 2); + +by (subgoal_tac "new_tv n2 s & new_tv n2 r & new_tv n2 ra" 1); +by (asm_full_simp_tac (!simpset addsimps [new_tv_subst]) 1); +by ((forward_tac [new_tv_subst_tel] 1) THEN (atac 1)); +by ((dres_inst_tac [("xc","$ s a")] new_tv_W 1) THEN (atac 1)); +by (safe_tac HOL_cs); +by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] + addss !simpset) 1); +by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] + addss !simpset) 1); +by (dres_inst_tac [("e","expr1")] (sym RS W_var_geD) 1); +by ((dtac new_tv_subst_tel 1) THEN (atac 1)); +by ((dres_inst_tac [("a","$ s a")] new_tv_list_le 1) THEN (atac 1)); +by ((dtac new_tv_subst_tel 1) THEN (atac 1)); +by (fast_tac (HOL_cs addDs [new_tv_W] addss + (!simpset addsimps [subst_comp_tel])) 1); + +qed "I_correct_wrt_W"; diff -r e74f694ca1da -r c7a8f374339b src/HOL/MiniML/I.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MiniML/I.thy Wed Oct 25 09:46:46 1995 +0100 @@ -0,0 +1,24 @@ +(* Title: MiniML.thy + Author: Thomas Stauner + Copyright 1995 TU Muenchen + + Recursive definition of type inference algorithm I for Mini_ML. +*) + +I = W + + +consts + I :: "[expr, type_expr list, nat, subst] => result_W" + +primrec I expr + I_Var "I (Var i) a n s = (if i < length a then Ok(s, nth i a, n) + else Fail)" + I_Abs "I (Abs e) a n s = I e ((TVar n)#a) (Suc n) s bind + (%(s,t,m). Ok(s, Fun (TVar n) t, m) )" + I_App "I (App e1 e2) a n s = + I e1 a n s bind + (%(s1,t1,m1). I e2 a m1 s1 bind + (%(s2,t2,m2). mgu ($s2 t1) (Fun ($s2 t2) (TVar m2)) bind + (%u. Ok($u o s2, TVar m2, Suc m2) ) ))" + +end diff -r e74f694ca1da -r c7a8f374339b src/HOL/MiniML/Maybe.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MiniML/Maybe.ML Wed Oct 25 09:46:46 1995 +0100 @@ -0,0 +1,20 @@ +open Maybe; + +(* constructor laws for bind *) +goalw thy [bind_def] "(Ok s) bind f = (f s)"; +by (Simp_tac 1); +qed "bind_Ok"; + +goalw thy [bind_def] "Fail bind f = Fail"; +by (Simp_tac 1); +qed "bind_Fail"; + +Addsimps [bind_Ok,bind_Fail]; + +(* expansion of bind *) +goal thy + "P(res bind f) = ((res = Fail --> P Fail) & (!s. res = Ok s --> P(f s)))"; +by (maybe.induct_tac "res" 1); +by (fast_tac (HOL_cs addss !simpset) 1); +by (Asm_simp_tac 1); +qed "expand_bind"; diff -r e74f694ca1da -r c7a8f374339b src/HOL/MiniML/Maybe.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MiniML/Maybe.thy Wed Oct 25 09:46:46 1995 +0100 @@ -0,0 +1,18 @@ +(* Title: HOL/MiniML/Maybe.thy + ID: $Id$ + Author: Dieter Nazareth and Tobias Nipkow + Copyright 1995 TU Muenchen + +Universal error monad. +*) + +Maybe = List + + +datatype 'a maybe = Ok 'a | Fail + +consts bind :: "['a maybe, 'a => 'b maybe] => 'b maybe" (infixl 60) + +defs + bind_def "m bind f == case m of Ok r => f r | Fail => Fail" + +end diff -r e74f694ca1da -r c7a8f374339b src/HOL/MiniML/MiniML.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MiniML/MiniML.ML Wed Oct 25 09:46:46 1995 +0100 @@ -0,0 +1,25 @@ +(* Title: HOL/MiniML/MiniML.ML + ID: $Id$ + Author: Dieter Nazareth and Tobias Nipkow + Copyright 1995 TU Muenchen +*) + +open MiniML; + +Addsimps has_type.intrs; +Addsimps [Un_upper1,Un_upper2]; + + +(* has_type is closed w.r.t. substitution *) +goal MiniML.thy + "!a e t. (a |- e :: t) --> ($ s a |- e :: $ s t)"; +by (rtac has_type.mutual_induct 1); +(* case VarI *) +by (asm_full_simp_tac (!simpset addsimps [app_subst_list]) 1); +by (forw_inst_tac [("f1","$ s")] (nth_map RS sym) 1); +by( fast_tac (HOL_cs addIs [has_type.VarI] addss (!simpset delsimps [nth_map])) 1); +(* case AbsI *) +by (asm_full_simp_tac (!simpset addsimps [app_subst_list]) 1); +(* case AppI *) +by (Asm_full_simp_tac 1); +bind_thm ("has_type_cl_sub", result() RS spec RS spec RS spec RS mp); diff -r e74f694ca1da -r c7a8f374339b src/HOL/MiniML/MiniML.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MiniML/MiniML.thy Wed Oct 25 09:46:46 1995 +0100 @@ -0,0 +1,32 @@ +(* Title: HOL/MiniML/MiniML.thy + ID: $Id$ + Author: Dieter Nazareth and Tobias Nipkow + Copyright 1995 TU Muenchen + +Mini_ML with type inference rules. +*) + +MiniML = Type + + +(* expressions *) +datatype + expr = Var nat | Abs expr | App expr expr + +(* type inference rules *) +consts + has_type :: "(type_expr list * expr * type_expr)set" +syntax + "@has_type" :: "[type_expr list, expr, type_expr] => bool" + ("((_) |-/ (_) :: (_))" 60) +translations + "a |- e :: t" == "(a,e,t):has_type" + +inductive "has_type" +intrs + VarI "[| n < length a |] ==> a |- Var n :: nth n a" + AbsI "[| t1#a |- e :: t2 |] ==> a |- Abs e :: Fun t1 t2" + AppI "[| a |- e1 :: Fun t2 t1; a |- e2 :: t2 |] + ==> a |- App e1 e2 :: t1" + +end + diff -r e74f694ca1da -r c7a8f374339b src/HOL/MiniML/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MiniML/ROOT.ML Wed Oct 25 09:46:46 1995 +0100 @@ -0,0 +1,17 @@ +(* Title: HOL/MiniML/ROOT.ML + ID: $Id$ + Author: Tobias Nipkow + Copyright 1995 TUM + +Type inference for let-free MiniML +*) + +HOL_build_completed; (*Make examples fail if HOL did*) + +writeln"Root file for HOL/MiniML"; +loadpath := [".","MiniML"]; +Unify.trace_bound := 20; + +time_use_thy "I"; + +make_chart (); (*make HTML chart*) diff -r e74f694ca1da -r c7a8f374339b src/HOL/MiniML/Type.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MiniML/Type.ML Wed Oct 25 09:46:46 1995 +0100 @@ -0,0 +1,341 @@ +open Type; + +Addsimps [app_subst_TVar,app_subst_Fun]; +Addsimps [mgu_eq,mgu_mg,mgu_free]; +Addsimps [free_tv_TVar,free_tv_Fun,free_tv_Nil,free_tv_Cons]; + +(* mgu does not introduce new type variables *) +goalw thy [new_tv_def] + "!! n. [|mgu t1 t2 = Ok u; new_tv n t1; new_tv n t2|] ==> \ +\ new_tv n u"; +by( fast_tac (set_cs addDs [mgu_free]) 1); +qed "mgu_new"; + +(* application of id_subst does not change type expression *) +goalw thy [id_subst_def] + "$ id_subst = (%t::type_expr.t)"; +by (rtac ext 1); +by (type_expr.induct_tac "t" 1); +by (ALLGOALS Asm_simp_tac); +qed "app_subst_id_te"; +Addsimps [app_subst_id_te]; + +(* application of id_subst does not change list of type expressions *) +goalw thy [app_subst_list] + "$ id_subst = (%l::type_expr list.l)"; +by (rtac ext 1); +by (list.induct_tac "l" 1); +by (ALLGOALS Asm_simp_tac); +qed "app_subst_id_tel"; +Addsimps [app_subst_id_tel]; + +goalw thy [dom_def,id_subst_def,empty_def] + "dom id_subst = {}"; +by (Simp_tac 1); +qed "dom_id_subst"; +Addsimps [dom_id_subst]; + +goalw thy [cod_def] + "cod id_subst = {}"; +by (Simp_tac 1); +qed "cod_id_subst"; +Addsimps [cod_id_subst]; + +goalw thy [free_tv_subst] + "free_tv id_subst = {}"; +by (Simp_tac 1); +qed "free_tv_id_subst"; +Addsimps [free_tv_id_subst]; + +goalw thy [dom_def,cod_def,UNION_def,Bex_def] + "!!v. [| v : free_tv(s n); v~=n |] ==> v : cod s"; +by (Simp_tac 1); +by (safe_tac (empty_cs addSIs [exI,conjI,notI])); +ba 2; +by (rotate_tac 1 1); +by (Asm_full_simp_tac 1); +qed "cod_app_subst"; +Addsimps [cod_app_subst]; + + +(* composition of substitutions *) +goal thy + "$ g ($ f t::type_expr) = $ (%x. $ g (f x) ) t"; +by (type_expr.induct_tac "t" 1); +by (ALLGOALS Asm_simp_tac); +qed "subst_comp_te"; + +goalw thy [app_subst_list] + "$ g ($ f l::type_expr list) = $ (%x. $ g (f x)) l"; +by (list.induct_tac "l" 1); +(* case [] *) +by (Simp_tac 1); +(* case x#xl *) +by (asm_full_simp_tac (!simpset addsimps [subst_comp_te]) 1); +qed "subst_comp_tel"; + + +(* constructor laws for app_subst *) +goalw thy [app_subst_list] + "$ s [] = []"; +by (Simp_tac 1); +qed "app_subst_Nil"; + +goalw thy [app_subst_list] + "$ s (x#l) = ($ s x)#($ s l)"; +by (Simp_tac 1); +qed "app_subst_Cons"; + +Addsimps [app_subst_Nil,app_subst_Cons]; + +(* constructor laws for new_tv *) +goalw thy [new_tv_def] + "new_tv n (TVar m) = (m (s m = TVar m)) & \ +\ (! l. l < n --> new_tv n (s l) ))"; +by( safe_tac HOL_cs ); +(* ==> *) +by( fast_tac (HOL_cs addDs [leD] addss (!simpset addsimps [free_tv_subst,dom_def])) 1); +by( subgoal_tac "m:cod s | s l = TVar l" 1); +by( safe_tac HOL_cs ); +by(fast_tac (HOL_cs addDs [UnI2] addss (!simpset addsimps [free_tv_subst])) 1); +by(dres_inst_tac [("P","%x. m:free_tv x")] subst 1 THEN atac 1); +by(Asm_full_simp_tac 1); +by(fast_tac (set_cs addss (!simpset addsimps [free_tv_subst,cod_def,dom_def])) 1); +(* <== *) +by( rewrite_goals_tac [free_tv_subst,cod_def,dom_def] ); +by( safe_tac set_cs ); +by( cut_inst_tac [("m","m"),("n","n")] less_linear 1); +by( fast_tac (HOL_cs addSIs [less_or_eq_imp_le]) 1); +by( cut_inst_tac [("m","ma"),("n","n")] less_linear 1); +by( fast_tac (HOL_cs addSIs [less_or_eq_imp_le]) 1); +qed "new_tv_subst"; + +goal thy + "new_tv n = list_all (new_tv n)"; +br ext 1; +by(list.induct_tac "x" 1); +by(ALLGOALS Asm_simp_tac); +qed "new_tv_list"; + +(* substitution affects only variables occurring freely *) +goal thy + "new_tv n (t::type_expr) --> $(%x. if x=n then t' else s x) t = $s t"; +by (type_expr.induct_tac "t" 1); +by (ALLGOALS(asm_simp_tac (!simpset setloop (split_tac [expand_if])))); +qed "subst_te_new_tv"; +Addsimps [subst_te_new_tv]; + +goal thy + "new_tv n (a::type_expr list) --> $(%x. if x=n then t else s x) a = $s a"; +by (list.induct_tac "a" 1); +by (ALLGOALS Asm_full_simp_tac); +qed "subst_tel_new_tv"; +Addsimps [subst_tel_new_tv]; + +(* all greater variables are also new *) +goal thy + "n<=m --> new_tv n (t::type_expr) --> new_tv m t"; +by (type_expr.induct_tac "t" 1); +(* case TVar n *) +by( fast_tac (HOL_cs addIs [less_le_trans] addss !simpset) 1); +(* case Fun t1 t2 *) +by (Asm_simp_tac 1); +bind_thm ("new_tv_le",result() RS mp RS mp); +Addsimps [lessI RS less_imp_le RS new_tv_le]; + +goal thy + "n<=m --> new_tv n (a::type_expr list) --> new_tv m a"; +by( list.induct_tac "a" 1); +(* case [] *) +by(Simp_tac 1); +(* case a#al *) +by (fast_tac (HOL_cs addIs [new_tv_le] addss !simpset) 1); +bind_thm ("new_tv_list_le",result() RS mp RS mp); +Addsimps [lessI RS less_imp_le RS new_tv_list_le]; + +goal thy + "!! n. [| n<=m; new_tv n (s::subst) |] ==> new_tv m s"; +by (asm_full_simp_tac (!simpset addsimps [new_tv_subst]) 1); +by (rtac conjI 1); +by (slow_tac (HOL_cs addIs [le_trans]) 1); +by (safe_tac HOL_cs ); +by (res_inst_tac [("P","l < n"),("Q","n <= l")] disjE 1); +by (fast_tac (HOL_cs addss (HOL_ss addsimps [le_def])) 1); +by (ALLGOALS (asm_full_simp_tac (!simpset addsimps [new_tv_le])) ); +qed "new_tv_subst_le"; +Addsimps [lessI RS less_imp_le RS new_tv_subst_le]; + +(* new_tv property remains if a substitution is applied *) +goal thy + "!!n. [| n new_tv m (s n)"; +by (asm_full_simp_tac (!simpset addsimps [new_tv_subst]) 1); +qed "new_tv_subst_var"; + +goal thy + "new_tv n s --> new_tv n (t::type_expr) --> new_tv n ($ s t)"; +by (type_expr.induct_tac "t" 1); +by (ALLGOALS(fast_tac (HOL_cs addss (!simpset addsimps [new_tv_subst])))); +bind_thm ("new_tv_subst_te",result() RS mp RS mp); +Addsimps [new_tv_subst_te]; + +goal thy + "new_tv n s --> new_tv n (a::type_expr list) --> new_tv n ($ s a)"; +by( simp_tac (!simpset addsimps [new_tv_list]) 1); +by (list.induct_tac "a" 1); +by (ALLGOALS(fast_tac (HOL_cs addss (!simpset addsimps [new_tv_subst])))); +bind_thm ("new_tv_subst_tel",result() RS mp RS mp); +Addsimps [new_tv_subst_tel]; + +(* auxilliary lemma *) +goal thy + "new_tv n a --> new_tv (Suc n) ((TVar n)#a)"; +by( simp_tac (!simpset addsimps [new_tv_list]) 1); +by (list.induct_tac "a" 1); +by (ALLGOALS Asm_full_simp_tac); +qed "new_tv_Suc_list"; + + +(* composition of substitutions preserves new_tv proposition *) +goal thy + "!! n. [| new_tv n (s::subst); new_tv n r |] ==> \ +\ new_tv n (($ r) o s)"; +by (asm_full_simp_tac (!simpset addsimps [new_tv_subst]) 1); +qed "new_tv_subst_comp_1"; + +goal thy + "!! n. [| new_tv n (s::subst); new_tv n r |] ==> \ +\ new_tv n (%v.$ r (s v))"; +by (asm_full_simp_tac (!simpset addsimps [new_tv_subst]) 1); +qed "new_tv_subst_comp_2"; + +Addsimps [new_tv_subst_comp_1,new_tv_subst_comp_2]; + +(* new type variables do not occur freely in a type structure *) +goalw thy [new_tv_def] + "!!n. new_tv n ts ==> n~:(free_tv ts)"; +by (fast_tac (HOL_cs addEs [less_anti_refl]) 1); +qed "new_tv_not_free_tv"; +Addsimps [new_tv_not_free_tv]; + +goal thy + "(t::type_expr) mem l --> free_tv t <= free_tv l"; +by (list.induct_tac "l" 1); +(* case [] *) +by (Simp_tac 1); +(* case x#xl *) +by (fast_tac (set_cs addss (!simpset setloop (split_tac [expand_if]))) 1); +bind_thm ("ftv_mem_sub_ftv_list",result() RS mp); +Addsimps [ftv_mem_sub_ftv_list]; + + +(* if two substitutions yield the same result if applied to a type + structure the substitutions coincide on the free type variables + occurring in the type structure *) +goal thy + "$ s1 (t::type_expr) = $ s2 t --> n:(free_tv t) --> s1 n = s2 n"; +by (type_expr.induct_tac "t" 1); +(* case TVar n *) +by (fast_tac (HOL_cs addss !simpset) 1); +(* case Fun t1 t2 *) +by (fast_tac (HOL_cs addss !simpset) 1); +bind_thm ("eq_subst_te_eq_free",result() RS mp RS mp); + +goal thy + "(!n. n:(free_tv t) --> s1 n = s2 n) --> $ s1 (t::type_expr) = $ s2 t"; +by (type_expr.induct_tac "t" 1); +(* case TVar n *) +by (fast_tac (HOL_cs addss !simpset) 1); +(* case Fun t1 t2 *) +by (fast_tac (HOL_cs addss !simpset) 1); +bind_thm ("eq_free_eq_subst_te",result() RS mp); + +goal thy + "$ s1 (l::type_expr list) = $ s2 l --> n:(free_tv l) --> s1 n = s2 n"; +by (list.induct_tac "l" 1); +(* case [] *) +by (fast_tac (HOL_cs addss !simpset) 1); +(* case x#xl *) +by (fast_tac (HOL_cs addIs [eq_subst_te_eq_free] addss (!simpset)) 1); +bind_thm ("eq_subst_tel_eq_free",result() RS mp RS mp); + +goal thy + "(!n. n:(free_tv l) --> s1 n = s2 n) --> $s1 (l::type_expr list) = $s2 l"; +by (list.induct_tac "l" 1); +(* case [] *) +by (fast_tac (HOL_cs addss !simpset) 1); +(* case x#xl *) +by (fast_tac (HOL_cs addIs [eq_free_eq_subst_te] addss (!simpset)) 1); +bind_thm ("eq_free_eq_subst_tel",result() RS mp); + +(* some useful theorems *) +goalw thy [free_tv_subst] + "!!v. v : cod s ==> v : free_tv s"; +by( fast_tac set_cs 1); +qed "codD"; + +goalw thy [free_tv_subst,dom_def] + "!! x. x ~: free_tv s ==> s x = TVar x"; +by( fast_tac set_cs 1); +qed "not_free_impl_id"; + +goalw thy [new_tv_def] + "!! n. [| new_tv n t; m:free_tv t |] ==> m 'a::type_struct" ("$") + +rules + app_subst_TVar "$ s (TVar n) = s n" + app_subst_Fun "$ s (Fun t1 t2) = Fun ($ s t1) ($ s t2)" +defs + app_subst_list "$ s == map ($ s)" + +(* free_tv s: the type variables occuring freely in the type structure s *) +consts + free_tv :: "['a::type_struct] => nat set" + +rules + free_tv_TVar "free_tv (TVar m) = {m}" + free_tv_Fun "free_tv (Fun t1 t2) = (free_tv t1) Un (free_tv t2)" + free_tv_Nil "free_tv [] = {}" + free_tv_Cons "free_tv (x#l) = (free_tv x) Un (free_tv l)" + +(* domain of a substitution *) +consts + dom :: "subst => nat set" +defs + dom_def "dom s == {n. s n ~= TVar n}" + +(* codomain of a substitutions: the introduced variables *) +consts + cod :: "subst => nat set" +defs + cod_def "cod s == (UN m:dom s. free_tv (s m))" + +defs + free_tv_subst "free_tv s == (dom s) Un (cod s)" + +(* new_tv s n computes whether n is a new type variable w.r.t. a type + structure s, i.e. whether n is greater than any type variable + occuring in the type structure *) +consts + new_tv :: "[nat,'a::type_struct] => bool" +defs + new_tv_def "new_tv n ts == ! m. m:free_tv ts --> m subst maybe" +rules + mgu_eq "mgu t1 t2 = Ok u ==> $ u t1 = $ u t2" + mgu_mg "[| (mgu t1 t2) = Ok u; $ s t1 = $ s t2 |] ==> + ? r. s = ($ r) o u" + mgu_Ok "$ s t1 = $ s t2 ==> ? u. mgu t1 t2 = Ok u" + mgu_free "mgu t1 t2 = Ok u ==> free_tv u <= free_tv t1 Un free_tv t2" + +end + diff -r e74f694ca1da -r c7a8f374339b src/HOL/MiniML/W.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MiniML/W.ML Wed Oct 25 09:46:46 1995 +0100 @@ -0,0 +1,379 @@ +(* Title: HOL/MiniML/W.ML + ID: $Id$ + Author: Dieter Nazareth and Tobias Nipkow + Copyright 1995 TU Muenchen + +Correctness and completeness of type inference algorithm W +*) + +open W; + + +(* stronger version of Suc_leD *) +goalw Nat.thy [le_def] + "!!m. Suc m <= n ==> m < n"; +by (Asm_full_simp_tac 1); +by (cut_facts_tac [less_linear] 1); +by (fast_tac HOL_cs 1); +qed "Suc_le_lessD"; +Addsimps [Suc_le_lessD]; + +(* correctness of W with respect to has_type *) +goal thy + "!a s t m n . Ok (s,t,m) = W e a n --> ($ s a |- e :: t)"; +by (expr.induct_tac "e" 1); +(* case Var n *) +by (asm_simp_tac (!simpset setloop (split_tac [expand_if])) 1); +(* case Abs e *) +by (asm_full_simp_tac (!simpset addsimps [app_subst_list] + setloop (split_tac [expand_bind])) 1); +by (strip_tac 1); +by (eres_inst_tac [("x","TVar(n) # a")] allE 1); +by( fast_tac (HOL_cs addss (!simpset addsimps [eq_sym_conv])) 1); +(* case App e1 e2 *) +by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1); +by (strip_tac 1); +by( rename_tac "sa ta na sb tb nb sc" 1); +by (res_inst_tac [("t2.0","$ sc tb")] has_type.AppI 1); +by (res_inst_tac [("s1","sc")] (app_subst_TVar RS subst) 1); +by (rtac (app_subst_Fun RS subst) 1); +by (res_inst_tac [("t","$ sc (Fun tb (TVar nb))"),("s","$ sc ($ sb ta)")] subst 1); +by (Asm_full_simp_tac 1); +by (simp_tac (HOL_ss addsimps [subst_comp_tel RS sym]) 1); +by ( (rtac has_type_cl_sub 1) THEN (rtac has_type_cl_sub 1)); +by (asm_full_simp_tac (!simpset addsimps [eq_sym_conv]) 1); +by (asm_full_simp_tac (!simpset addsimps [subst_comp_tel RS sym,has_type_cl_sub,eq_sym_conv]) 1); +qed "W_correct"; + +val has_type_casesE = map(has_type.mk_cases expr.simps) + [" s |- Var n :: t"," s |- Abs e :: t","s |- App e1 e2 ::t"]; + + +(* the resulting type variable is always greater or equal than the given one *) +goal thy + "!a n s t m. W e a n = Ok (s,t,m) --> n<=m"; +by (expr.induct_tac "e" 1); +(* case Var(n) *) +by (fast_tac (HOL_cs addss (!simpset setloop (split_tac [expand_if]))) 1); +(* case Abs e *) +by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1); +by (fast_tac (HOL_cs addDs [Suc_leD]) 1); +(* case App e1 e2 *) +by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1); +by (strip_tac 1); +by (rename_tac "s t na sa ta nb sb sc tb m" 1); +by (eres_inst_tac [("x","a")] allE 1); +by (eres_inst_tac [("x","n")] allE 1); +by (eres_inst_tac [("x","$ s a")] allE 1); +by (eres_inst_tac [("x","s")] allE 1); +by (eres_inst_tac [("x","t")] allE 1); +by (eres_inst_tac [("x","na")] allE 1); +by (eres_inst_tac [("x","na")] allE 1); +by (asm_full_simp_tac (!simpset addsimps [eq_sym_conv]) 1); +by (etac conjE 1); +by (eres_inst_tac [("x","sa")] allE 1); +by (eres_inst_tac [("x","ta")] allE 1); +by (eres_inst_tac [("x","nb")] allE 1); +by (etac conjE 1); +by (res_inst_tac [("j","na")] le_trans 1); +by (Asm_simp_tac 1); +by (asm_simp_tac (!simpset addsimps [Suc_leD]) 1); +qed "W_var_ge"; + +Addsimps [W_var_ge]; + +goal thy + "!! s. Ok (s,t,m) = W e a n ==> n<=m"; +by (asm_full_simp_tac (!simpset addsimps [eq_sym_conv]) 1); +qed "W_var_geD"; + + +(* auxiliary lemma *) +goal Maybe.thy "(y = Ok x) = (Ok x = y)"; +by( simp_tac (!simpset addsimps [eq_sym_conv]) 1); +qed "rotate_Ok"; + + +(* resulting type variable is new *) +goal thy + "!n a s t m. new_tv n a --> W e a n = Ok (s,t,m) --> \ +\ (new_tv m s & new_tv m t)"; +by (expr.induct_tac "e" 1); +(* case Var n *) +by (fast_tac (HOL_cs addss (!simpset + addsimps [id_subst_def,list_all_nth,new_tv_list,new_tv_subst] + setloop (split_tac [expand_if]))) 1); + +(* case Abs e *) +by (simp_tac (!simpset addsimps [new_tv_subst,new_tv_Suc_list] + setloop (split_tac [expand_bind])) 1); +by (strip_tac 1); +by (eres_inst_tac [("x","Suc n")] allE 1); +by (eres_inst_tac [("x","(TVar n)#a")] allE 1); +by (fast_tac (HOL_cs addss (!simpset + addsimps [new_tv_subst,new_tv_Suc_list])) 1); + +(* case App e1 e2 *) +by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1); +by (strip_tac 1); +by (rename_tac "s t na sa ta nb sb sc tb m" 1); +by (eres_inst_tac [("x","n")] allE 1); +by (eres_inst_tac [("x","a")] allE 1); +by (eres_inst_tac [("x","s")] allE 1); +by (eres_inst_tac [("x","t")] allE 1); +by (eres_inst_tac [("x","na")] allE 1); +by (eres_inst_tac [("x","na")] allE 1); +by (asm_full_simp_tac (!simpset addsimps [eq_sym_conv]) 1); +by (eres_inst_tac [("x","$ s a")] allE 1); +by (eres_inst_tac [("x","sa")] allE 1); +by (eres_inst_tac [("x","ta")] allE 1); +by (eres_inst_tac [("x","nb")] allE 1); +by( asm_full_simp_tac (!simpset addsimps [rotate_Ok]) 1); +by (rtac conjI 1); +by (rtac new_tv_subst_comp_2 1); +by (rtac new_tv_subst_comp_2 1); +by (rtac (lessI RS less_imp_le RS new_tv_subst_le) 1); +by (res_inst_tac [("n","na")] new_tv_subst_le 1); +by (asm_full_simp_tac (!simpset addsimps [rotate_Ok]) 1); +by (Asm_simp_tac 1); +by (fast_tac (HOL_cs addDs [W_var_geD] addIs + [new_tv_list_le,new_tv_subst_tel,lessI RS less_imp_le RS new_tv_subst_le]) + 1); +be (sym RS mgu_new) 1; +by (fast_tac (HOL_cs addDs [W_var_geD] addIs [new_tv_subst_te,new_tv_list_le, + new_tv_subst_tel,lessI RS less_imp_le RS new_tv_le,lessI RS less_imp_le RS + new_tv_subst_le,new_tv_le]) 1); +by (fast_tac (HOL_cs addDs [W_var_geD] addIs + [new_tv_list_le,new_tv_subst_tel,new_tv_le] + addss (!simpset)) 1); +br (lessI RS new_tv_subst_var) 1; +be (sym RS mgu_new) 1; +by (fast_tac (HOL_cs addSIs [lessI RS less_imp_le RS new_tv_le,new_tv_subst_te] + addDs [W_var_geD] addIs + [new_tv_list_le,new_tv_subst_tel,lessI RS less_imp_le RS + new_tv_subst_le,new_tv_le] addss !simpset) 1); +by (fast_tac (HOL_cs addDs [W_var_geD] addIs + [new_tv_list_le,new_tv_subst_tel,new_tv_le] + addss (!simpset)) 1); +bind_thm ("new_tv_W",result() RS spec RS spec RS spec RS spec RS spec RS mp RS mp); + + +goal thy + "!n a s t m v. W e a n = Ok (s,t,m) --> \ +\ (v:free_tv s | v:free_tv t) --> v v:free_tv a"; +by (expr.induct_tac "e" 1); +(* case Var n *) +by (fast_tac (HOL_cs addIs [nth_mem,subsetD,ftv_mem_sub_ftv_list] + addss (!simpset setloop (split_tac [expand_if]))) 1); + +(* case Abs e *) +by (asm_full_simp_tac (!simpset addsimps + [free_tv_subst] setloop (split_tac [expand_bind])) 1); +by (strip_tac 1); +by (rename_tac "s t na sa ta m v" 1); +by (eres_inst_tac [("x","Suc n")] allE 1); +by (eres_inst_tac [("x","TVar n # a")] allE 1); +by (eres_inst_tac [("x","s")] allE 1); +by (eres_inst_tac [("x","t")] allE 1); +by (eres_inst_tac [("x","na")] allE 1); +by (eres_inst_tac [("x","v")] allE 1); +by (fast_tac (HOL_cs addIs [cod_app_subst] addss !simpset) 1); + +(* case App e1 e2 *) +by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1); +by (strip_tac 1); +by (rename_tac "s t na sa ta nb sb sc tb m v" 1); +by (eres_inst_tac [("x","n")] allE 1); +by (eres_inst_tac [("x","a")] allE 1); +by (eres_inst_tac [("x","s")] allE 1); +by (eres_inst_tac [("x","t")] allE 1); +by (eres_inst_tac [("x","na")] allE 1); +by (eres_inst_tac [("x","na")] allE 1); +by (eres_inst_tac [("x","v")] allE 1); +(* second case *) +by (eres_inst_tac [("x","$ s a")] allE 1); +by (eres_inst_tac [("x","sa")] allE 1); +by (eres_inst_tac [("x","ta")] allE 1); +by (eres_inst_tac [("x","nb")] allE 1); +by (eres_inst_tac [("x","v")] allE 1); +by (safe_tac (empty_cs addSIs [conjI,impI] addSEs [conjE]) ); +by (asm_full_simp_tac (!simpset addsimps [rotate_Ok]) 1); +bd W_var_geD 1; +bd W_var_geD 1; +by ( (forward_tac [less_le_trans] 1) THEN (assume_tac 1) ); +by (fast_tac (HOL_cs addDs [free_tv_comp_subst RS subsetD,sym RS mgu_free, + codD,free_tv_app_subst_te RS subsetD,free_tv_app_subst_tel RS subsetD, + less_le_trans,less_not_refl2,subsetD] + addEs [UnE] + addss !simpset) 1); +by (Asm_full_simp_tac 1); +bd (sym RS W_var_geD) 1; +bd (sym RS W_var_geD) 1; +by ( (forward_tac [less_le_trans] 1) THEN (assume_tac 1) ); +by (fast_tac (HOL_cs addDs [mgu_free, codD,free_tv_subst_var RS subsetD, + free_tv_app_subst_te RS subsetD,free_tv_app_subst_tel RS subsetD, + less_le_trans,subsetD] + addEs [UnE] + addss !simpset) 1); +bind_thm ("free_tv_W",result() RS spec RS spec RS spec RS spec RS spec RS spec RS mp RS mp RS mp); + +goal HOL.thy "(~(P | Q)) = (~P & ~Q)"; +by( fast_tac HOL_cs 1); +qed "not_disj"; + +(* Completeness of W w.r.t. has_type *) +goal thy + "!s' a t' n. ($ s' a |- e :: t') --> new_tv n a --> \ +\ (? s t. (? m. W e a n = Ok (s,t,m) ) & \ +\ (? r. $ s' a = $ r ($ s a) & \ +\ t' = $ r t ) )"; +by (expr.induct_tac "e" 1); +(* case Var n *) +by (strip_tac 1); +by (simp_tac (!simpset addcongs [conj_cong] + setloop (split_tac [expand_if])) 1); +by (eresolve_tac has_type_casesE 1); +by (asm_full_simp_tac (!simpset addsimps [eq_sym_conv,app_subst_list]) 1); +by (res_inst_tac [("x","id_subst")] exI 1); +by (res_inst_tac [("x","nth nat a")] exI 1); +by (Simp_tac 1); +by (res_inst_tac [("x","s'")] exI 1); +by (Asm_simp_tac 1); + +(* case Abs e *) +by (strip_tac 1); +by (eresolve_tac has_type_casesE 1); +by (eres_inst_tac [("x","%x.if x=n then t1 else (s' x)")] allE 1); +by (eres_inst_tac [("x","(TVar n)#a")] allE 1); +by (eres_inst_tac [("x","t2")] allE 1); +by (eres_inst_tac [("x","Suc n")] allE 1); +by (fast_tac (HOL_cs addss (!simpset addcongs [conj_cong] + setloop (split_tac [expand_bind]))) 1); + +(* case App e1 e2 *) +by (strip_tac 1); +by (eresolve_tac has_type_casesE 1); +by (eres_inst_tac [("x","s'")] allE 1); +by (eres_inst_tac [("x","a")] allE 1); +by (eres_inst_tac [("x","Fun t2 t'")] allE 1); +by (eres_inst_tac [("x","n")] allE 1); +by (safe_tac HOL_cs); +by (eres_inst_tac [("x","r")] allE 1); +by (eres_inst_tac [("x","$ s a")] allE 1); +by (eres_inst_tac [("x","t2")] allE 1); +by (eres_inst_tac [("x","m")] allE 1); +bd asm_rl 1; +bd asm_rl 1; +bd asm_rl 1; +by (Asm_full_simp_tac 1); +by (safe_tac HOL_cs); +by (fast_tac HOL_cs 1); +by (fast_tac (HOL_cs addIs [sym RS W_var_geD,new_tv_W RS + conjunct1,new_tv_list_le,new_tv_subst_tel]) 1); + +by (subgoal_tac + "$ (%x.if x=ma then t' else (if x:(free_tv t - free_tv sa) then r x \ +\ else ra x)) ($ sa t) = \ +\ $ (%x.if x=ma then t' else (if x:(free_tv t - free_tv sa) then r x \ +\ else ra x)) (Fun ta (TVar ma))" 1); +by (res_inst_tac [("t","$ (%x. if x = ma then t' else \ +\ (if x:(free_tv t - free_tv sa) then r x else ra x)) ($ sa t)"), + ("s","Fun ($ ra ta) t'")] ssubst 2); +by (asm_simp_tac (!simpset addsimps [subst_comp_te]) 2); +br eq_free_eq_subst_te 2; +by (strip_tac 2); +by (subgoal_tac "na ~=ma" 2); +by (fast_tac (HOL_cs addDs [new_tv_W,sym RS W_var_geD, + new_tv_not_free_tv,new_tv_le]) 3); +by (case_tac "na:free_tv sa" 2); +(* na ~: free_tv sa *) +by (asm_simp_tac (!simpset addsimps [not_free_impl_id] + setloop (split_tac [expand_if])) 3); +(* na : free_tv sa *) +by (dres_inst_tac [("l1","$ s a")] (subst_comp_tel RSN (2,trans)) 2); +bd eq_subst_tel_eq_free 2; +by (fast_tac (HOL_cs addIs [free_tv_W,free_tv_le_new_tv] addDs [new_tv_W]) 2); +by (Asm_simp_tac 2); +by (case_tac "na:dom sa" 2); +(* na ~: dom sa *) +by (asm_full_simp_tac (!simpset addsimps [dom_def] + setloop (split_tac [expand_if])) 3); +(* na : dom sa *) +br eq_free_eq_subst_te 2; +by (strip_tac 2); +by (subgoal_tac "nb ~= ma" 2); +by ((forward_tac [new_tv_W] 3) THEN (atac 3)); +be conjE 3; +bd new_tv_subst_tel 3; +by (fast_tac (HOL_cs addIs [new_tv_list_le] addDs [sym RS W_var_geD]) 3); +by (fast_tac (set_cs addDs [new_tv_W,new_tv_not_free_tv] addss + (!simpset addsimps [cod_def,free_tv_subst])) 3); +by (fast_tac (set_cs addss (!simpset addsimps [cod_def,free_tv_subst] + setloop (split_tac [expand_if]))) 2); + +by (Simp_tac 2); +br eq_free_eq_subst_te 2; +by (strip_tac 2 ); +by (subgoal_tac "na ~= ma" 2); +by ((forward_tac [new_tv_W] 3) THEN (atac 3)); +be conjE 3; +bd (sym RS W_var_geD) 3; +by (fast_tac (HOL_cs addDs [new_tv_list_le,new_tv_subst_tel,new_tv_W,new_tv_not_free_tv]) 3); +by (case_tac "na: free_tv t - free_tv sa" 2); +(* case na ~: free_tv t - free_tv sa *) +by( asm_full_simp_tac (!simpset setloop (split_tac [expand_if])) 3); +(* case na : free_tv t - free_tv sa *) +by( asm_full_simp_tac (!simpset setloop (split_tac [expand_if])) 2); +by (dres_inst_tac [("l1","$ s a")] (subst_comp_tel RSN (2,trans)) 2); +bd eq_subst_tel_eq_free 2; +by (fast_tac (HOL_cs addIs [free_tv_W,free_tv_le_new_tv] addDs [new_tv_W]) 2); +by (asm_full_simp_tac (!simpset addsimps [free_tv_subst,dom_def,not_disj]) 2); + +by (asm_simp_tac (!simpset setloop (split_tac [expand_bind])) 1); +by (safe_tac HOL_cs ); +bd mgu_Ok 1; +by( fast_tac (HOL_cs addss !simpset) 1); +by (REPEAT (resolve_tac [exI,conjI] 1)); +by (fast_tac HOL_cs 1); +by (fast_tac HOL_cs 1); +by ((dtac mgu_mg 1) THEN (atac 1)); +be exE 1; +by (res_inst_tac [("x","rb")] exI 1); +br conjI 1; +by (dres_inst_tac [("x","ma")] fun_cong 2); +by( asm_full_simp_tac (!simpset addsimps [eq_sym_conv]) 2); +by (simp_tac (!simpset addsimps [subst_comp_tel RS sym]) 1); +by (res_inst_tac [("l2","($ sa ($ s a))")] ((subst_comp_tel RS sym) RSN + (2,trans)) 1); +by( asm_full_simp_tac (!simpset addsimps [o_def,eq_sym_conv]) 1); +br eq_free_eq_subst_tel 1; +by( safe_tac HOL_cs ); +by (subgoal_tac "ma ~= na" 1); +by ((forward_tac [new_tv_W] 2) THEN (atac 2)); +be conjE 2; +bd new_tv_subst_tel 2; +by (fast_tac (HOL_cs addIs [new_tv_list_le] addDs [sym RS W_var_geD]) 2); +by (( forw_inst_tac [("xd","m")] (sym RSN (2,new_tv_W)) 2) THEN (atac 2)); +be conjE 2; +bd (free_tv_app_subst_tel RS subsetD) 2; +by (fast_tac (set_cs addDs [W_var_geD,new_tv_list_le,codD, + new_tv_not_free_tv]) 2); +by (case_tac "na: free_tv t - free_tv sa" 1); +(* case na ~: free_tv t - free_tv sa *) +by (asm_full_simp_tac (!simpset setloop (split_tac [expand_if])) 2); +(* case na : free_tv t - free_tv sa *) +by (asm_full_simp_tac (!simpset setloop (split_tac [expand_if])) 1); +bd (free_tv_app_subst_tel RS subsetD) 1; +by (fast_tac (set_cs addDs [codD,subst_comp_tel RSN (2,trans), + eq_subst_tel_eq_free] addss ((!simpset addsimps + [not_disj,free_tv_subst,dom_def]))) 1); +qed "W_complete"; + + + + + + + + + diff -r e74f694ca1da -r c7a8f374339b src/HOL/MiniML/W.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MiniML/W.thy Wed Oct 25 09:46:46 1995 +0100 @@ -0,0 +1,28 @@ +(* Title: HOL/MiniML/W.thy + ID: $Id$ + Author: Dieter Nazareth and Tobias Nipkow + Copyright 1995 TU Muenchen + +Type inference algorithm W +*) + +W = MiniML + + +types + result_W = "(subst * type_expr * nat)maybe" + +(* type inference algorithm W *) +consts + W :: "[expr, type_expr list, nat] => result_W" + +primrec W expr + W_Var "W (Var i) a n = (if i < length a then Ok(id_subst, nth i a, n) + else Fail)" + W_Abs "W (Abs e) a n = W e ((TVar n)#a) (Suc n) bind + (%(s,t,m). Ok(s, Fun (s n) t, m) )" + W_App "W (App e1 e2) a n = + W e1 a n bind + (%(s1,t1,m1). W e2 ($ s1 a) m1 bind + (%(s2,t2,m2). mgu ($ s2 t1) (Fun t2 (TVar m2)) bind + (%u. Ok( ($ u) o (($ s2) o s1), $ u (TVar m2), Suc m2) )))" +end