diff -r ed54ec416bb3 -r 5c816239deaa Unification/Substs.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Unification/Substs.thy Sun Apr 29 11:29:56 2012 +0100 @@ -0,0 +1,363 @@ + + +theory Substs = Main + Terms + PreEqu + Equ: + +(* substitutions *) + +types substs = "(string \ trm)list" + +consts + look_up :: "string \ substs \ trm" +primrec + "look_up X [] = Susp [] X" + "look_up X (x#xs) = (if (X = fst x) then (snd x) else look_up X xs)" + +consts + subst :: "substs \ trm \ trm" +primrec + subst_unit: "subst s (Unit) = Unit" + subst_susp: "subst s (Susp pi X) = swap pi (look_up X s)" + subst_atom: "subst s (Atom a) = Atom a" + subst_abst: "subst s (Abst a t) = Abst a (subst s t)" + subst_paar: "subst s (Paar t1 t2) = Paar (subst s t1) (subst s t2)" + subst_func: "subst s (Func F t) = Func F (subst s t)" + +declare subst_susp [simp del] + +(* composition of substitutions (adapted from Martin Coen) *) + +consts + alist_rec :: "substs \ substs \ (string\trm\substs\substs\substs) \ substs" + +primrec + "alist_rec [] c d = c" + "alist_rec (p#al) c d = d (fst p) (snd p) al (alist_rec al c d)" + +consts + "\" :: "substs \ substs \ substs" (infixl 81) +defs + comp_def: "s1 \ s2 \ alist_rec s2 s1 (\ x y xs g. (x,subst s1 y)#g)" + +(* domain of substitutions *) + +consts + domn :: "(trm \ trm) \ string set" +defs + domn_def: "domn s \ {X. (s (Susp [] X)) \ (Susp [] X)}" + +(* substitutions extending freshness environments *) + +consts + ext_subst :: "fresh_envs \ (trm \ trm) \ fresh_envs \ bool" (" _ \ _ _ " [80,80,80] 80) +defs + ext_subst_def: "nabla1 \ s (nabla2) \ (\(a,X)\nabla2. nabla1\a\ s (Susp [] X))" + +(* alpah-equality for substitutions *) + +consts + subst_equ :: "fresh_envs \ (trm\trm) \ (trm\trm) \ bool" (" _ \ _ \ _" [80,80,80] 80) + +defs + subst_equ_def: + "nabla\ s1 \ s2 \ \X\(domn s1\domn s2). (nabla \ s1 (Susp [] X) \ s2 (Susp [] X))" + +lemma not_in_domn: "X\(domn s)\ (s (Susp [] X)) = (Susp [] X)" +apply(auto simp only: domn_def) +done + +lemma subst_swap_comm: "subst s (swap pi t) = swap pi (subst s t)" +apply(induct_tac t) +apply(auto simp add: swap_append subst_susp) +done + +lemma subst_not_occurs: "\(occurs X t) \ subst [(X,t2)] t = t" +apply(induct t) +apply(auto simp add: subst_susp) +done + +lemma [simp]: "subst [] t = t" +apply(induct_tac t, auto simp add: subst_susp) +done + +lemma [simp]: "subst (s\[]) = subst s" +apply(auto simp add: comp_def) +done + +lemma [simp]: "subst ([]\s) = subst s" +apply(rule ext) +apply(induct_tac x) +apply(auto) +apply(induct_tac s rule: alist_rec.induct) +apply(auto simp add: comp_def subst_susp) +done + +lemma subst_comp_expand: "subst (s1\s2) t = subst s1 (subst s2 t)" +apply(induct_tac t) +apply(auto) +apply(induct_tac s2 rule: alist_rec.induct) +apply(auto simp add: comp_def subst_susp subst_swap_comm) +done + +lemma subst_assoc: "subst (s1\(s2\s3))=subst ((s1\s2)\s3)" +apply(rule ext) +apply(induct_tac x) +apply(auto) +apply(simp add: subst_comp_expand) +done + +lemma fresh_subst: "nabla1\a\t\ nabla2\(subst s) nabla1 \ nabla2\a\subst s t" +apply(erule fresh.induct) +apply(auto) +--Susp +apply(simp add: ext_subst_def subst_susp) +apply(drule_tac x="(swapas (rev pi) a, X)" in bspec) +apply(assumption) +apply(simp) +apply(rule fresh_swap_right[THEN mp]) +apply(assumption) +done + +lemma equ_subst: "nabla1\t1\t2\ nabla2\ (subst s) nabla1 \ nabla2\(subst s t1)\(subst s t2)" +apply(erule equ.induct) +apply(auto) +apply(rule_tac "nabla1.1"="nabla" in fresh_subst[THEN mp]) +apply(assumption)+ +apply(simp add: subst_swap_comm) +-- Susp +apply(simp only: subst_susp) +apply(rule equ_pi1_pi2_add[THEN mp]) +apply(simp only: ext_subst_def subst_susp) +apply(force) +done + +lemma unif_1: + "\nabla\subst s (Susp pi X)\subst s t\\ nabla\ subst (s\[(X,swap (rev pi) t)])\subst s" +apply(simp only: subst_equ_def) +apply(case_tac "pi=[]") +apply(simp add: subst_susp comp_def) +apply(force intro: equ_sym equ_refl) +apply(subgoal_tac "domn (subst (s\[(X,swap (rev pi) t)]))=domn(subst s)\{X}")--A +apply(simp) +apply(rule conjI) +apply(simp add: subst_comp_expand) +apply(simp add: subst_susp) +apply(simp only: subst_swap_comm) +apply(simp only: equ_pi_to_left) +apply(rule equ_sym) +apply(assumption) +apply(rule ballI) +apply(simp only: subst_comp_expand) +apply(simp add: subst_susp) +apply(force intro: equ_sym equ_refl simp add: subst_swap_comm equ_pi_to_left) +--A +apply(simp only: domn_def) +apply(auto) +apply(simp add: subst_comp_expand) +apply(simp add: subst_susp subst_swap_comm) +apply(simp only: subst_comp_expand) +apply(simp add: subst_susp subst_swap_comm) +apply(drule swap_empty[THEN mp]) +apply(simp) +apply(simp only: subst_comp_expand) +apply(simp only: subst_susp) +apply(auto) +apply(case_tac "x=X") +apply(simp) +apply(simp add: subst_swap_comm) +apply(drule swap_empty[THEN mp]) +apply(simp) +apply(simp add: subst_susp) +done + +lemma subst_equ_a: +"\nabla\(subst s1) \ (subst s2)\\ \t2. nabla\(subst s2 t1)\t2 \ nabla\(subst s1 t1)\t2" +apply(rule allI) +apply(induct t1) +--Abst.ab +apply(rule impI) +apply(simp) +apply(ind_cases "nabla \ Abst list (subst s1 trm) \ t2") +apply(best) +--Abst.aa +apply(force) +--Susp +apply(rule impI) +apply(simp only: subst_equ_def) +apply(case_tac "list2\domn (subst s1) \ domn (subst s2)") +apply(drule_tac x="list2" in bspec) +apply(assumption) +apply(simp) +apply(subgoal_tac "nabla \ subst s2 (Susp [] list2) \ swap (rev list1) t2")--A +apply(drule_tac "t1.0"="subst s1 (Susp [] list2)" and + "t2.0"="subst s2 (Susp [] list2)" and + "t3.0"="swap (rev list1) t2" in equ_trans) +apply(assumption) +apply(simp only: equ_pi_to_right) +apply(simp add: subst_swap_comm[THEN sym]) +--A +apply(simp only: equ_pi_to_right) +apply(simp add: subst_swap_comm[THEN sym]) +apply(simp) +apply(erule conjE) +apply(drule not_in_domn)+ +apply(subgoal_tac "subst s1 (Susp list1 list2)=swap list1 (subst s1 (Susp [] list2))")--B +apply(subgoal_tac "subst s2 (Susp list1 list2)=swap list1 (subst s2 (Susp [] list2))")--C +apply(simp) +--BC +apply(simp (no_asm) add: subst_swap_comm[THEN sym]) +apply(simp (no_asm) add: subst_swap_comm[THEN sym]) +--Unit +apply(force) +--Atom +apply(force) +--Paar +apply(rule impI) +apply(simp) +apply(ind_cases "nabla \ Paar (subst sigma1 trm1) (subst sigma1 trm2) \ t2") +apply(best) +--Func +apply(rule impI) +apply(simp) +apply(ind_cases "nabla \ Func list (subst sigma1 trm) \ t2") +apply(best) +done + +lemma unif_2a: "\nabla\subst s1\subst s2\\ + (nabla\subst s2 t1 \ subst s2 t2)\(nabla\subst s1 t1 \ subst s1 t2)" +apply(rule impI) +apply(frule_tac "t1.0"="t1" in subst_equ_a) +apply(drule_tac x="subst s2 t2" in spec) +apply(simp) +apply(drule equ_sym) +apply(drule equ_sym) +apply(frule_tac "t1.0"="t2" in subst_equ_a) +apply(drule_tac x="subst s1 t1" in spec) +apply(rule equ_sym) +apply(simp) +done + + +lemma unif_2b: "\nabla\subst s1\ subst s2\\nabla\a\ subst s2 t\nabla\a\subst s1 t" +apply(induct t) +--Abst +apply(simp) +apply(rule impI) +apply(case_tac "a=list") +apply(force) +apply(force dest!: fresh_abst_ab_elim) +--Susp +apply(rule impI) +apply(subgoal_tac "subst s1 (Susp list1 list2)= swap list1 (subst s1 (Susp [] list2))")--A +apply(subgoal_tac "subst s2 (Susp list1 list2)= swap list1 (subst s2 (Susp [] list2))")--B +apply(simp add: subst_equ_def) +apply(case_tac "list2\domn (subst s1) \ domn (subst s2)") +apply(drule_tac x="list2" in bspec) +apply(assumption) +apply(rule fresh_swap_right[THEN mp]) +apply(rule_tac "t1.1"=" subst s2 (Susp [] list2)" in l3_jud[THEN mp]) +apply(rule equ_sym) +apply(simp) +apply(rule fresh_swap_left[THEN mp]) +apply(simp) +apply(simp) +apply(erule conjE) +apply(drule not_in_domn)+ +apply(simp add: subst_swap_comm) +--AB +apply(simp (no_asm) add: subst_swap_comm[THEN sym]) +apply(simp (no_asm) add: subst_swap_comm[THEN sym]) +--Unit +apply(force) +--Atom +apply(force) +--Paar +apply(force dest!: fresh_paar_elim) +--Func +apply(force dest!: fresh_func_elim) +done + +lemma subst_equ_to_trm: "nabla\subst s1 \ subst s2\ nabla\subst s1 t\subst s2 t" +apply(induct_tac t) +apply(auto simp add: subst_equ_def) +apply(case_tac "list2\domn (subst s1) \ domn (subst s2)") +apply(simp only: subst_susp) +apply(simp only: equ_pi_to_left[THEN sym]) +apply(simp only: equ_involutive_left) +apply(simp) +apply(erule conjE) +apply(drule not_in_domn)+ +apply(subgoal_tac "subst s1 (Susp list1 list2)=swap list1 (subst s1 (Susp [] list2))")--A +apply(subgoal_tac "subst s2 (Susp list1 list2)=swap list1 (subst s2 (Susp [] list2))")--B +apply(simp) +apply(rule equ_refl) +apply(simp (no_asm_use) add: subst_swap_comm[THEN sym])+ +done + +lemma subst_cancel_right: "\nabla\(subst s1)\(subst s2)\\nabla\(subst (s1\s))\(subst (s2\s))" +apply(simp (no_asm) only: subst_equ_def) +apply(rule ballI) +apply(simp only: subst_comp_expand) +apply(rule subst_equ_to_trm) +apply(assumption) +done + +lemma subst_trans: "\nabla\subst s1\subst s2; nabla\subst s2\subst s3\\nabla\subst s1\subst s3" +apply(simp add: subst_equ_def) +apply(auto) +apply(case_tac "X \domn (subst s2)") +apply(drule_tac x="X" in bspec) +apply(force) +apply(drule_tac x="X" in bspec) +apply(force) +apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans) +apply(assumption)+ +apply(case_tac "X \domn (subst s3)") +apply(drule_tac x="X" in bspec) +apply(force) +apply(drule_tac x="X" in bspec) +apply(force) +apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans) +apply(assumption)+ +apply(drule not_in_domn)+ +apply(drule_tac x="X" in bspec) +apply(force) +apply(simp) +apply(case_tac "X \domn (subst s1)") +apply(drule_tac x="X" in bspec) +apply(force) +apply(drule_tac x="X" in bspec) +apply(force) +apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans) +apply(assumption)+ +apply(case_tac "X \domn (subst s2)") +apply(drule_tac x="X" in bspec) +apply(force) +apply(drule_tac x="X" in bspec) +apply(force) +apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans) +apply(assumption)+ +apply(drule not_in_domn)+ +apply(simp) +apply(rotate_tac 1) +apply(drule_tac x="X" in bspec) +apply(force) +apply(simp) +done + +(* if occurs holds, then one subterm is equal to (subst s (Susp pi X)) *) + +lemma occurs_sub_trm_equ: + "occurs X t1 \ (\t2\sub_trms (subst s t1). (\pi.(nabla\subst s (Susp pi1 X)\(swap pi t2))))" +apply(induct_tac t1) +apply(auto) +apply(simp only: subst_susp) +apply(rule_tac x="swap list1 (look_up X s)" in bexI) +apply(rule_tac x="pi1@(rev list1)" in exI) +apply(simp add: swap_append) +apply(simp add: equ_pi_to_left[THEN sym]) +apply(simp only: equ_involutive_left) +apply(rule equ_refl) +apply(rule t_sub_trms_t) +done + +end