# HG changeset patch # User wenzelm # Date 1188293129 -7200 # Node ID 448978b61556422ace84c36e724f6cc33053a6af # Parent ab6206ccb570b24e139f76564abc2b6b70b70ab4 induct: proper separation of initial and terminal step; avoid unspecific prems; diff -r ab6206ccb570 -r 448978b61556 src/HOL/ex/Unification.thy --- a/src/HOL/ex/Unification.thy Tue Aug 28 03:58:37 2007 +0200 +++ b/src/HOL/ex/Unification.thy Tue Aug 28 11:25:29 2007 +0200 @@ -178,7 +178,7 @@ show "\ =\<^sub>s [(v,t)] \ \" proof fix s show "s \ \ = s \ [(v,t)] \ \" using th - by (induct s, auto) + by (induct s) auto qed qed @@ -194,7 +194,7 @@ assumes "unify_dom (M, N)" assumes "unify M N = Some \" shows "MGU \ M N" -using prems +using assms proof (induct M N arbitrary: \) case (7 M N M' N' \) -- "The interesting case" @@ -357,21 +357,21 @@ assumes "unify M N = Some \" shows "vars_of (t \ \) \ vars_of M \ vars_of N \ vars_of t" (is "?P M N \ t") -using prems +using assms proof (induct M N arbitrary:\ t) case (3 c v) hence "\ = [(v, Const c)]" by simp - thus ?case by (induct t, auto) + thus ?case by (induct t) auto next case (4 M N v) hence "\occ (Var v) (M\N)" by (cases "occ (Var v) (M\N)", auto) - with prems have "\ = [(v, M\N)]" by simp - thus ?case by (induct t, auto) + with 4 have "\ = [(v, M\N)]" by simp + thus ?case by (induct t) auto next case (5 v M) hence "\occ (Var v) M" by (cases "occ (Var v) M", auto) - with prems have "\ = [(v, M)]" by simp - thus ?case by (induct t, auto) + with 5 have "\ = [(v, M)]" by simp + thus ?case by (induct t) auto next case (7 M N M' N' \) then obtain \1 \2 @@ -420,7 +420,7 @@ assumes "unify M N = Some \" shows "(\v\vars_of M \ vars_of N. elim \ v) \ \ =\<^sub>s []" (is "?P M N \") -using prems +using assms proof (induct M N arbitrary:\) case 1 thus ?case by simp next @@ -428,19 +428,19 @@ next case (3 c v) have no_occ: "\ occ (Var v) (Const c)" by simp - with prems have "\ = [(v, Const c)]" by simp + with 3 have "\ = [(v, Const c)]" by simp with occ_elim[OF no_occ] show ?case by auto next case (4 M N v) hence no_occ: "\occ (Var v) (M\N)" by (cases "occ (Var v) (M\N)", auto) - with prems have "\ = [(v, M\N)]" by simp + with 4 have "\ = [(v, M\N)]" by simp with occ_elim[OF no_occ] show ?case by auto next case (5 v M) hence no_occ: "\occ (Var v) M" by (cases "occ (Var v) M", auto) - with prems have "\ = [(v, M)]" by simp + with 5 have "\ = [(v, M)]" by simp with occ_elim[OF no_occ] show ?case by auto next