# HG changeset patch # User krauss # Date 1313957584 -7200 # Node ID 74c08021ab2e741643cec4e07cf69ccc6953187f # Parent 75ec83d45303dbb57c5682c4e02a3d8002c5f94c changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone. diff -r 75ec83d45303 -r 74c08021ab2e src/HOL/Import/Generate-HOL/GenHOL4Base.thy --- a/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Sun Aug 21 11:03:15 2011 -0700 +++ b/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Sun Aug 21 22:13:04 2011 +0200 @@ -28,11 +28,12 @@ TYPE_DEFINITION > HOL4Setup.TYPE_DEFINITION LET > HOL4Compat.LET; -ignore_thms +(*ignore_thms BOUNDED_DEF BOUNDED_THM UNBOUNDED_DEF UNBOUNDED_THM; +*) end_import; diff -r 75ec83d45303 -r 74c08021ab2e src/HOL/Import/Generate-HOL/ROOT.ML --- a/src/HOL/Import/Generate-HOL/ROOT.ML Sun Aug 21 11:03:15 2011 -0700 +++ b/src/HOL/Import/Generate-HOL/ROOT.ML Sun Aug 21 22:13:04 2011 +0200 @@ -1,3 +1,5 @@ +Runtime.debug := true; + use_thy "GenHOL4Prob"; use_thy "GenHOL4Vec"; use_thy "GenHOL4Word32"; diff -r 75ec83d45303 -r 74c08021ab2e src/HOL/Subst/AList.thy --- a/src/HOL/Subst/AList.thy Sun Aug 21 11:03:15 2011 -0700 +++ b/src/HOL/Subst/AList.thy Sun Aug 21 22:13:04 2011 +0200 @@ -23,4 +23,6 @@ "P [] \ (!!x y xs. P xs \ P ((x,y) # xs)) \ P l" by (induct l) auto + + end diff -r 75ec83d45303 -r 74c08021ab2e src/HOL/ex/Unification.thy --- a/src/HOL/ex/Unification.thy Sun Aug 21 11:03:15 2011 -0700 +++ b/src/HOL/ex/Unification.thy Sun Aug 21 22:13:04 2011 +0200 @@ -29,7 +29,7 @@ datatype 'a trm = Var 'a | Const 'a - | App "'a trm" "'a trm" (infix "\" 60) + | Comb "'a trm" "'a trm" (infix "\" 60) type_synonym 'a subst = "('a \ 'a trm) list" @@ -42,38 +42,38 @@ text {* Applying a substitution to a term: *} -fun apply_subst :: "'a trm \ 'a subst \ 'a trm" (infixl "\" 60) +primrec subst :: "'a trm \ 'a subst \ 'a trm" (infixl "\" 55) where - "(Var v) \ s = assoc v (Var v) s" -| "(Const c) \ s = (Const c)" -| "(M \ N) \ s = (M \ s) \ (N \ s)" + "(Var v) \ s = assoc v (Var v) s" +| "(Const c) \ s = (Const c)" +| "(M \ N) \ s = (M \ s) \ (N \ s)" text {* Composition of substitutions: *} fun - "compose" :: "'a subst \ 'a subst \ 'a subst" (infixl "\" 80) + comp :: "'a subst \ 'a subst \ 'a subst" (infixl "\" 56) where - "[] \ bl = bl" -| "((a,b) # al) \ bl = (a, b \ bl) # (al \ bl)" + "[] \ bl = bl" +| "((a,b) # al) \ bl = (a, b \ bl) # (al \ bl)" text {* Equivalence of substitutions: *} -definition eqv (infix "=\<^sub>s" 50) +definition subst_eq (infixr "\" 52) where - "s1 =\<^sub>s s2 \ \t. t \ s1 = t \ s2" + "s1 \ s2 \ \t. t \ s1 = t \ s2" subsection {* Basic lemmas *} -lemma apply_empty[simp]: "t \ [] = t" +lemma apply_empty[simp]: "t \ [] = t" by (induct t) auto -lemma compose_empty[simp]: "\ \ [] = \" +lemma compose_empty[simp]: "\ \ [] = \" by (induct \) auto -lemma apply_compose[simp]: "t \ (s1 \ s2) = t \ s1 \ s2" +lemma apply_compose[simp]: "t \ (s1 \ s2) = t \ s1 \ s2" proof (induct t) - case App thus ?case by simp + case Comb thus ?case by simp next case Const thus ?case by simp next @@ -85,39 +85,39 @@ qed qed -lemma eqv_refl[intro]: "s =\<^sub>s s" - by (auto simp:eqv_def) +lemma eqv_refl[intro]: "s \ s" + by (auto simp:subst_eq_def) -lemma eqv_trans[trans]: "\s1 =\<^sub>s s2; s2 =\<^sub>s s3\ \ s1 =\<^sub>s s3" - by (auto simp:eqv_def) +lemma eqv_trans[trans]: "\s1 \ s2; s2 \ s3\ \ s1 \ s3" + by (auto simp:subst_eq_def) -lemma eqv_sym[sym]: "\s1 =\<^sub>s s2\ \ s2 =\<^sub>s s1" - by (auto simp:eqv_def) +lemma eqv_sym[sym]: "\s1 \ s2\ \ s2 \ s1" + by (auto simp:subst_eq_def) -lemma eqv_intro[intro]: "(\t. t \ \ = t \ \) \ \ =\<^sub>s \" - by (auto simp:eqv_def) +lemma eqv_intro[intro]: "(\t. t \ \ = t \ \) \ \ \ \" + by (auto simp:subst_eq_def) -lemma eqv_dest[dest]: "s1 =\<^sub>s s2 \ t \ s1 = t \ s2" - by (auto simp:eqv_def) +lemma eqv_dest[dest]: "s1 \ s2 \ t \ s1 = t \ s2" + by (auto simp:subst_eq_def) -lemma compose_eqv: "\\ =\<^sub>s \'; \ =\<^sub>s \'\ \ (\ \ \) =\<^sub>s (\' \ \')" - by (auto simp:eqv_def) +lemma compose_eqv: "\\ \ \'; \ \ \'\ \ (\ \ \) \ (\' \ \')" + by (auto simp:subst_eq_def) -lemma compose_assoc: "(a \ b) \ c =\<^sub>s a \ (b \ c)" +lemma compose_assoc: "(a \ b) \ c \ a \ (b \ c)" by auto subsection {* Specification: Most general unifiers *} definition - "Unifier \ t u \ (t\\ = u\\)" + "Unifier \ t u \ (t\\ = u\\)" definition "MGU \ t u \ Unifier \ t u \ (\\. Unifier \ t u - \ (\\. \ =\<^sub>s \ \ \))" + \ (\\. \ \ \ \ \))" lemma MGUI[intro]: - "\t \ \ = u \ \; \\. t \ \ = u \ \ \ \\. \ =\<^sub>s \ \ \\ + "\t \ \ = u \ \; \\. t \ \ = u \ \ \ \\. \ \ \ \ \\ \ MGU \ t u" by (simp only:Unifier_def MGU_def, auto) @@ -130,11 +130,11 @@ text {* Occurs check: Proper subterm relation *} -fun occ :: "'a trm \ 'a trm \ bool" +fun occs :: "'a trm \ 'a trm \ bool" (infixl "<:" 54) where - "occ u (Var v) = False" -| "occ u (Const c) = False" -| "occ u (M \ N) = (u = M \ u = N \ occ u M \ occ u N)" + "occs u (Var v) = False" +| "occs u (Const c) = False" +| "occs u (M \ N) = (u = M \ u = N \ occs u M \ occs u N)" text {* The unification algorithm: *} @@ -143,41 +143,41 @@ "unify (Const c) (M \ N) = None" | "unify (M \ N) (Const c) = None" | "unify (Const c) (Var v) = Some [(v, Const c)]" -| "unify (M \ N) (Var v) = (if (occ (Var v) (M \ N)) +| "unify (M \ N) (Var v) = (if (occs (Var v) (M \ N)) then None else Some [(v, M \ N)])" -| "unify (Var v) M = (if (occ (Var v) M) +| "unify (Var v) M = (if (occs (Var v) M) then None else Some [(v, M)])" | "unify (Const c) (Const d) = (if c=d then Some [] else None)" | "unify (M \ N) (M' \ N') = (case unify M M' of None \ None | - Some \ \ (case unify (N \ \) (N' \ \) + Some \ \ (case unify (N \ \) (N' \ \) of None \ None | - Some \ \ Some (\ \ \)))" + Some \ \ Some (\ \ \)))" by pat_completeness auto declare unify.psimps[simp] subsection {* Partial correctness *} -text {* Some lemmas about occ and MGU: *} +text {* Some lemmas about occs and MGU: *} -lemma subst_no_occ: "\occ (Var v) t \ Var v \ t - \ t \ [(v,s)] = t" +lemma subst_no_occs: "\occs (Var v) t \ Var v \ t + \ t \ [(v,s)] = t" by (induct t) auto lemma MGU_Var[intro]: - assumes no_occ: "\occ (Var v) t" + assumes no_occs: "\occs (Var v) t" shows "MGU [(v,t)] (Var v) t" proof (intro MGUI exI) - show "Var v \ [(v,t)] = t \ [(v,t)]" using no_occ - by (cases "Var v = t", auto simp:subst_no_occ) + show "Var v \ [(v,t)] = t \ [(v,t)]" using no_occs + by (cases "Var v = t", auto simp:subst_no_occs) next - fix \ assume th: "Var v \ \ = t \ \" - show "\ =\<^sub>s [(v,t)] \ \" + fix \ assume th: "Var v \ \ = t \ \" + show "\ \ [(v,t)] \ \" proof - fix s show "s \ \ = s \ [(v,t)] \ \" using th + fix s show "s \ \ = s \ [(v,t)] \ \" using th by (induct s) auto qed qed @@ -200,42 +200,42 @@ then obtain \1 \2 where "unify M M' = Some \1" - and "unify (N \ \1) (N' \ \1) = Some \2" - and \: "\ = \1 \ \2" + and "unify (N \ \1) (N' \ \1) = Some \2" + and \: "\ = \1 \ \2" and MGU_inner: "MGU \1 M M'" - and MGU_outer: "MGU \2 (N \ \1) (N' \ \1)" + and MGU_outer: "MGU \2 (N \ \1) (N' \ \1)" by (auto split:option.split_asm) show ?case proof from MGU_inner and MGU_outer - have "M \ \1 = M' \ \1" - and "N \ \1 \ \2 = N' \ \1 \ \2" + have "M \ \1 = M' \ \1" + and "N \ \1 \ \2 = N' \ \1 \ \2" unfolding MGU_def Unifier_def by auto - thus "M \ N \ \ = M' \ N' \ \" unfolding \ + thus "M \ N \ \ = M' \ N' \ \" unfolding \ by simp next - fix \' assume "M \ N \ \' = M' \ N' \ \'" - hence "M \ \' = M' \ \'" - and Ns: "N \ \' = N' \ \'" by auto + fix \' assume "M \ N \ \' = M' \ N' \ \'" + hence "M \ \' = M' \ \'" + and Ns: "N \ \' = N' \ \'" by auto with MGU_inner obtain \ - where eqv: "\' =\<^sub>s \1 \ \" + where eqv: "\' \ \1 \ \" unfolding MGU_def Unifier_def by auto - from Ns have "N \ \1 \ \ = N' \ \1 \ \" + from Ns have "N \ \1 \ \ = N' \ \1 \ \" by (simp add:eqv_dest[OF eqv]) with MGU_outer obtain \ - where eqv2: "\ =\<^sub>s \2 \ \" + where eqv2: "\ \ \2 \ \" unfolding MGU_def Unifier_def by auto - have "\' =\<^sub>s \ \ \" unfolding \ + have "\' \ \ \ \" unfolding \ by (rule eqv_intro, auto simp:eqv_dest[OF eqv] eqv_dest[OF eqv2]) - thus "\\. \' =\<^sub>s \ \ \" .. + thus "\\. \' \ \ \ \" .. qed qed (auto split:split_if_asm) -- "Solve the remaining cases automatically" @@ -256,50 +256,50 @@ text {* Elimination of variables by a substitution: *} definition - "elim \ v \ \t. v \ vars_of (t \ \)" + "elim \ v \ \t. v \ vars_of (t \ \)" -lemma elim_intro[intro]: "(\t. v \ vars_of (t \ \)) \ elim \ v" +lemma elim_intro[intro]: "(\t. v \ vars_of (t \ \)) \ elim \ v" by (auto simp:elim_def) -lemma elim_dest[dest]: "elim \ v \ v \ vars_of (t \ \)" +lemma elim_dest[dest]: "elim \ v \ v \ vars_of (t \ \)" by (auto simp:elim_def) -lemma elim_eqv: "\ =\<^sub>s \ \ elim \ x = elim \ x" - by (auto simp:elim_def eqv_def) +lemma elim_eqv: "\ \ \ \ elim \ x = elim \ x" + by (auto simp:elim_def subst_eq_def) text {* Replacing a variable by itself yields an identity subtitution: *} -lemma var_self[intro]: "[(v, Var v)] =\<^sub>s []" +lemma var_self[intro]: "[(v, Var v)] \ []" proof - fix t show "t \ [(v, Var v)] = t \ []" + fix t show "t \ [(v, Var v)] = t \ []" by (induct t) simp_all qed -lemma var_same: "([(v, t)] =\<^sub>s []) = (t = Var v)" +lemma var_same: "([(v, t)] \ []) = (t = Var v)" proof assume t_v: "t = Var v" - thus "[(v, t)] =\<^sub>s []" + thus "[(v, t)] \ []" by auto next - assume id: "[(v, t)] =\<^sub>s []" + assume id: "[(v, t)] \ []" show "t = Var v" proof - - have "t = Var v \ [(v, t)]" by simp - also from id have "\ = Var v \ []" .. + have "t = Var v \ [(v, t)]" by simp + also from id have "\ = Var v \ []" .. finally show ?thesis by simp qed qed -text {* A lemma about occ and elim *} +text {* A lemma about occs and elim *} lemma remove_var: assumes [simp]: "v \ vars_of s" - shows "v \ vars_of (t \ [(v, s)])" + shows "v \ vars_of (t \ [(v, s)])" by (induct t) simp_all -lemma occ_elim: "\occ (Var v) t - \ elim [(v,t)] v \ [(v,t)] =\<^sub>s []" +lemma occs_elim: "\occs (Var v) t + \ elim [(v,t)] v \ [(v,t)] \ []" proof (induct t) case (Var x) show ?case @@ -319,29 +319,29 @@ by (auto intro!:remove_var) thus ?case .. next - case (App M N) + case (Comb M N) - hence ih1: "elim [(v, M)] v \ [(v, M)] =\<^sub>s []" - and ih2: "elim [(v, N)] v \ [(v, N)] =\<^sub>s []" - and nonocc: "Var v \ M" "Var v \ N" + hence ih1: "elim [(v, M)] v \ [(v, M)] \ []" + and ih2: "elim [(v, N)] v \ [(v, N)] \ []" + and nonoccs: "Var v \ M" "Var v \ N" by auto - from nonocc have "\ [(v,M)] =\<^sub>s []" + from nonoccs have "\ [(v,M)] \ []" by (simp add:var_same) with ih1 have "elim [(v, M)] v" by blast - hence "v \ vars_of (Var v \ [(v,M)])" .. + hence "v \ vars_of (Var v \ [(v,M)])" .. hence not_in_M: "v \ vars_of M" by simp - from nonocc have "\ [(v,N)] =\<^sub>s []" + from nonoccs have "\ [(v,N)] \ []" by (simp add:var_same) with ih2 have "elim [(v, N)] v" by blast - hence "v \ vars_of (Var v \ [(v,N)])" .. + hence "v \ vars_of (Var v \ [(v,N)])" .. hence not_in_N: "v \ vars_of N" by simp have "elim [(v, M \ N)] v" proof fix t - show "v \ vars_of (t \ [(v, M \ N)])" + show "v \ vars_of (t \ [(v, M \ N)])" proof (induct t) case (Var x) thus ?case by (simp add: not_in_M not_in_N) qed auto @@ -354,7 +354,7 @@ lemma unify_vars: assumes "unify_dom (M, N)" assumes "unify M N = Some \" - shows "vars_of (t \ \) \ vars_of M \ vars_of N \ vars_of t" + shows "vars_of (t \ \) \ vars_of M \ vars_of N \ vars_of t" (is "?P M N \ t") using assms proof (induct M N arbitrary:\ t) @@ -363,45 +363,45 @@ 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) + hence "\occs (Var v) (M\N)" by (cases "occs (Var v) (M\N)", 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) + hence "\occs (Var v) M" by (cases "occs (Var v) M", 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 where "unify M M' = Some \1" - and "unify (N \ \1) (N' \ \1) = Some \2" - and \: "\ = \1 \ \2" + and "unify (N \ \1) (N' \ \1) = Some \2" + and \: "\ = \1 \ \2" and ih1: "\t. ?P M M' \1 t" - and ih2: "\t. ?P (N\\1) (N'\\1) \2 t" + and ih2: "\t. ?P (N\\1) (N'\\1) \2 t" by (auto split:option.split_asm) show ?case proof - fix v assume a: "v \ vars_of (t \ \)" + fix v assume a: "v \ vars_of (t \ \)" show "v \ vars_of (M \ N) \ vars_of (M' \ N') \ vars_of t" proof (cases "v \ vars_of M \ v \ vars_of M' \ v \ vars_of N \ v \ vars_of N'") case True - with ih1 have l:"\t. v \ vars_of (t \ \1) \ v \ vars_of t" + with ih1 have l:"\t. v \ vars_of (t \ \1) \ v \ vars_of t" by auto - from a and ih2[where t="t \ \1"] - have "v \ vars_of (N \ \1) \ vars_of (N' \ \1) - \ v \ vars_of (t \ \1)" unfolding \ + from a and ih2[where t="t \ \1"] + have "v \ vars_of (N \ \1) \ vars_of (N' \ \1) + \ v \ vars_of (t \ \1)" unfolding \ by auto hence "v \ vars_of t" proof - assume "v \ vars_of (N \ \1) \ vars_of (N' \ \1)" + assume "v \ vars_of (N \ \1) \ vars_of (N' \ \1)" with True show ?thesis by (auto dest:l) next - assume "v \ vars_of (t \ \1)" + assume "v \ vars_of (t \ \1)" thus ?thesis by (rule l) qed @@ -417,7 +417,7 @@ lemma unify_eliminates: assumes "unify_dom (M, N)" assumes "unify M N = Some \" - shows "(\v\vars_of M \ vars_of N. elim \ v) \ \ =\<^sub>s []" + shows "(\v\vars_of M \ vars_of N. elim \ v) \ \ \ []" (is "?P M N \") using assms proof (induct M N arbitrary:\) @@ -426,21 +426,21 @@ case 2 thus ?case by simp next case (3 c v) - have no_occ: "\ occ (Var v) (Const c)" by simp + have no_occs: "\ occs (Var v) (Const c)" by simp with 3 have "\ = [(v, Const c)]" by simp - with occ_elim[OF no_occ] + with occs_elim[OF no_occs] 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) + hence no_occs: "\occs (Var v) (M\N)" by (cases "occs (Var v) (M\N)", auto) with 4 have "\ = [(v, M\N)]" by simp - with occ_elim[OF no_occ] + with occs_elim[OF no_occs] show ?case by auto next case (5 v M) - hence no_occ: "\occ (Var v) M" by (cases "occ (Var v) M", auto) + hence no_occs: "\occs (Var v) M" by (cases "occs (Var v) M", auto) with 5 have "\ = [(v, M)]" by simp - with occ_elim[OF no_occ] + with occs_elim[OF no_occs] show ?case by auto next case (6 c d) thus ?case @@ -449,43 +449,43 @@ case (7 M N M' N' \) then obtain \1 \2 where "unify M M' = Some \1" - and "unify (N \ \1) (N' \ \1) = Some \2" - and \: "\ = \1 \ \2" + and "unify (N \ \1) (N' \ \1) = Some \2" + and \: "\ = \1 \ \2" and ih1: "?P M M' \1" - and ih2: "?P (N\\1) (N'\\1) \2" + and ih2: "?P (N\\1) (N'\\1) \2" by (auto split:option.split_asm) from `unify_dom (M \ N, M' \ N')` have "unify_dom (M, M')" by (rule accp_downward) (rule unify_rel.intros) hence no_new_vars: - "\t. vars_of (t \ \1) \ vars_of M \ vars_of M' \ vars_of t" + "\t. vars_of (t \ \1) \ vars_of M \ vars_of M' \ vars_of t" by (rule unify_vars) (rule `unify M M' = Some \1`) from ih2 show ?case proof - assume "\v\vars_of (N \ \1) \ vars_of (N' \ \1). elim \2 v" + assume "\v\vars_of (N \ \1) \ vars_of (N' \ \1). elim \2 v" then obtain v - where "v\vars_of (N \ \1) \ vars_of (N' \ \1)" + where "v\vars_of (N \ \1) \ vars_of (N' \ \1)" and el: "elim \2 v" by auto with no_new_vars show ?thesis unfolding \ by (auto simp:elim_def) next - assume empty[simp]: "\2 =\<^sub>s []" + assume empty[simp]: "\2 \ []" - have "\ =\<^sub>s (\1 \ [])" unfolding \ + have "\ \ (\1 \ [])" unfolding \ by (rule compose_eqv) auto - also have "\ =\<^sub>s \1" by auto - finally have "\ =\<^sub>s \1" . + also have "\ \ \1" by auto + finally have "\ \ \1" . from ih1 show ?thesis proof assume "\v\vars_of M \ vars_of M'. elim \1 v" - with elim_eqv[OF `\ =\<^sub>s \1`] + with elim_eqv[OF `\ \ \1`] show ?thesis by auto next - note `\ =\<^sub>s \1` - also assume "\1 =\<^sub>s []" + note `\ \ \1` + also assume "\1 \ []" finally show ?thesis .. qed qed @@ -509,7 +509,7 @@ "unify M M' = Some \" from unify_eliminates[OF inner] - show "((N \ \, N' \ \), (M \ N, M' \ N')) \?R" + show "((N \ \, N' \ \), (M \ N, M' \ N')) \?R" proof -- {* Either a variable is eliminated \ldots *} assume "(\v\vars_of M \ vars_of M'. elim \ v)" @@ -517,7 +517,7 @@ where "elim \ v" and "v\vars_of M \ vars_of M'" by auto with unify_vars[OF inner] - have "vars_of (N\\) \ vars_of (N'\\) + have "vars_of (N\\) \ vars_of (N'\\) \ vars_of (M\N) \ vars_of (M'\N')" by auto @@ -525,9 +525,9 @@ by (auto intro!: measures_less intro: psubset_card_mono) next -- {* Or the substitution is empty *} - assume "\ =\<^sub>s []" - hence "N \ \ = N" - and "N' \ \ = N'" by auto + assume "\ \ []" + hence "N \ \ = N" + and "N' \ \ = N'" by auto thus ?thesis by (auto intro!: measures_less intro: psubset_card_mono) qed