# HG changeset patch # User krauss # Date 1313957584 -7200 # Node ID 3a10392fb8c309f199a8fd5e3c4d0ba51b9d7f1e # Parent 03d91bfad83bd2f8a7716edb9ee76e86a7301a5b added proof of idempotence, roughly after HOL/Subst/Unify.thy diff -r 03d91bfad83b -r 3a10392fb8c3 src/HOL/ex/Unification.thy --- a/src/HOL/ex/Unification.thy Sun Aug 21 22:13:04 2011 +0200 +++ b/src/HOL/ex/Unification.thy Sun Aug 21 22:13:04 2011 +0200 @@ -118,6 +118,9 @@ lemma subst_trans[trans]: "\s1 \ s2; s2 \ s3\ \ s1 \ s3" by (auto simp:subst_eq_def) +lemma subst_no_occs: "\ Var v \ t \ Var v \ t + \ t \ [(v,s)] = t" +by (induct t) auto text {* Composition of Substitutions *} @@ -170,6 +173,9 @@ "MGU \ s t \ MGU \ t s" by (auto simp:MGU_def Unifier_def) +lemma MGU_is_Unifier: "MGU \ t u \ Unifier \ t u" +unfolding MGU_def by (rule conjunct1) + lemma MGU_Var: assumes "\ Var v \ t" shows "MGU [(v,t)] (Var v) t" @@ -431,7 +437,6 @@ lemma unify_computes_MGU: "unify M N = Some \ \ MGU \ M N" -using assms proof (induct M N arbitrary: \ rule: unify.induct) case (7 M N M' N' \) -- "The interesting case" @@ -476,6 +481,62 @@ qed qed (auto simp: MGU_Const intro: MGU_Var MGU_Var[symmetric] split: split_if_asm) +lemma Idem_Nil [iff]: "Idem []" + by (simp add: Idem_def) +lemma Var_Idem: + assumes "~ (Var v \ t)" shows "Idem [(v,t)]" + unfolding Idem_def +proof + from assms have [simp]: "t \ [(v, t)] = t" + by (metis assoc.simps(2) subst.simps(1) subst_no_occs) + + fix s show "s \ [(v, t)] \ [(v, t)] = s \ [(v, t)]" + by (induct s) auto +qed + +lemma Unifier_Idem_subst: + "Idem(r) \ Unifier s (t \ r) (u \ r) \ + Unifier (r \ s) (t \ r) (u \ r)" +by (simp add: Idem_def Unifier_def subst_eq_def) + +lemma Idem_comp: + "Idem r \ Unifier s (t \ r) (u \ r) \ + (!!q. Unifier q (t \ r) (u \ r) \ s \ q \ q) \ + Idem (r \ s)" + apply (frule Unifier_Idem_subst, blast) + apply (force simp add: Idem_def subst_eq_def) + done + +theorem unify_gives_Idem: + "unify M N = Some \ \ Idem \" +proof (induct M N arbitrary: \ rule: unify.induct) + case (7 M M' N N' \) + + then obtain \1 \2 + where "unify M N = Some \1" + and \2: "unify (M' \ \1) (N' \ \1) = Some \2" + and \: "\ = \1 \ \2" + and "Idem \1" + and "Idem \2" + by (auto split: option.split_asm) + + from \2 have "Unifier \2 (M' \ \1) (N' \ \1)" + by (rule unify_computes_MGU[THEN MGU_is_Unifier]) + + with `Idem \1` + show "Idem \" unfolding \ + proof (rule Idem_comp) + fix \ assume "Unifier \ (M' \ \1) (N' \ \1)" + with \2 obtain \ where \: "\ \ \2 \ \" + using unify_computes_MGU MGU_def by blast + + have "\2 \ \ \ \2 \ (\2 \ \)" by (rule subst_cong) (auto simp: \) + also have "... \ (\2 \ \2) \ \" by (rule comp_assoc[symmetric]) + also have "... \ \2 \ \" by (rule subst_cong) (auto simp: `Idem \2`[unfolded Idem_def]) + also have "... \ \" by (rule \[symmetric]) + finally show "\2 \ \ \ \" . + qed +qed (auto intro!: Var_Idem split: option.splits if_splits) end