Unification/Substs.thy
changeset 107 5c816239deaa
equal deleted inserted replaced
106:ed54ec416bb3 107:5c816239deaa
       
     1 
       
     2 
       
     3 theory Substs = Main + Terms + PreEqu + Equ:
       
     4 
       
     5 (* substitutions *)
       
     6 
       
     7 types substs = "(string \<times> trm)list"
       
     8 
       
     9 consts 
       
    10   look_up :: "string \<Rightarrow> substs \<Rightarrow> trm"
       
    11 primrec 
       
    12   "look_up X []     = Susp [] X"
       
    13   "look_up X (x#xs) = (if (X = fst x) then (snd x) else look_up X xs)"
       
    14 
       
    15 consts 
       
    16   subst :: "substs \<Rightarrow> trm \<Rightarrow> trm"
       
    17 primrec
       
    18   subst_unit: "subst s (Unit)          = Unit"
       
    19   subst_susp: "subst s (Susp pi X)     = swap pi (look_up X s)"
       
    20   subst_atom: "subst s (Atom a)        = Atom a" 
       
    21   subst_abst: "subst s (Abst a t)      = Abst a (subst s t)"
       
    22   subst_paar: "subst s (Paar t1 t2)    = Paar (subst s t1) (subst s t2)"
       
    23   subst_func: "subst s (Func F t)      = Func F (subst s t)"
       
    24 
       
    25 declare subst_susp [simp del]
       
    26 
       
    27 (* composition of substitutions (adapted from Martin Coen) *)
       
    28 
       
    29 consts
       
    30   alist_rec :: "substs \<Rightarrow> substs \<Rightarrow> (string\<Rightarrow>trm\<Rightarrow>substs\<Rightarrow>substs\<Rightarrow>substs) \<Rightarrow> substs"
       
    31 
       
    32 primrec
       
    33   "alist_rec [] c d = c"
       
    34   "alist_rec (p#al) c d = d (fst p) (snd p) al (alist_rec al c d)"
       
    35 
       
    36 consts 
       
    37   "\<bullet>"   ::  "substs \<Rightarrow> substs \<Rightarrow> substs" (infixl 81)
       
    38 defs
       
    39   comp_def: "s1 \<bullet> s2 \<equiv> alist_rec s2 s1 (\<lambda> x y xs g. (x,subst s1 y)#g)"
       
    40 
       
    41 (* domain of substitutions *)
       
    42 
       
    43 consts  
       
    44   domn :: "(trm \<Rightarrow> trm) \<Rightarrow> string set"
       
    45 defs 
       
    46   domn_def: "domn s \<equiv> {X. (s (Susp [] X)) \<noteq> (Susp [] X)}" 
       
    47 
       
    48 (* substitutions extending freshness environments *)
       
    49 
       
    50 consts 
       
    51   ext_subst :: "fresh_envs \<Rightarrow> (trm \<Rightarrow> trm) \<Rightarrow> fresh_envs \<Rightarrow> bool" (" _ \<Turnstile> _ _ " [80,80,80] 80)
       
    52 defs 
       
    53   ext_subst_def: "nabla1 \<Turnstile> s (nabla2) \<equiv> (\<forall>(a,X)\<in>nabla2. nabla1\<turnstile>a\<sharp> s (Susp [] X))"
       
    54 
       
    55 (* alpah-equality for substitutions *)
       
    56 
       
    57 consts 
       
    58   subst_equ :: "fresh_envs \<Rightarrow> (trm\<Rightarrow>trm) \<Rightarrow> (trm\<Rightarrow>trm) \<Rightarrow> bool" (" _ \<Turnstile> _ \<approx> _" [80,80,80] 80)
       
    59 
       
    60 defs 
       
    61   subst_equ_def: 
       
    62   "nabla\<Turnstile> s1 \<approx> s2 \<equiv>  \<forall>X\<in>(domn s1\<union>domn s2). (nabla \<turnstile> s1 (Susp [] X) \<approx> s2 (Susp [] X))"
       
    63 
       
    64 lemma not_in_domn: "X\<notin>(domn s)\<Longrightarrow> (s (Susp [] X)) = (Susp [] X)"
       
    65 apply(auto simp only: domn_def)
       
    66 done
       
    67 
       
    68 lemma subst_swap_comm: "subst s (swap pi t) = swap pi (subst s t)"
       
    69 apply(induct_tac t)
       
    70 apply(auto simp add: swap_append subst_susp)
       
    71 done
       
    72 
       
    73 lemma subst_not_occurs: "\<not>(occurs X t) \<longrightarrow> subst [(X,t2)] t = t"
       
    74 apply(induct t)
       
    75 apply(auto simp add: subst_susp)
       
    76 done
       
    77 
       
    78 lemma [simp]: "subst [] t = t"
       
    79 apply(induct_tac t, auto simp add: subst_susp)
       
    80 done
       
    81 
       
    82 lemma [simp]: "subst (s\<bullet>[]) = subst s"
       
    83 apply(auto simp add: comp_def)
       
    84 done
       
    85 
       
    86 lemma [simp]: "subst ([]\<bullet>s) = subst s"
       
    87 apply(rule ext)
       
    88 apply(induct_tac x)
       
    89 apply(auto)
       
    90 apply(induct_tac s rule: alist_rec.induct)
       
    91 apply(auto simp add: comp_def subst_susp)
       
    92 done
       
    93 
       
    94 lemma subst_comp_expand: "subst (s1\<bullet>s2) t = subst s1 (subst s2 t)"
       
    95 apply(induct_tac t)
       
    96 apply(auto)
       
    97 apply(induct_tac s2 rule: alist_rec.induct)
       
    98 apply(auto simp add: comp_def subst_susp subst_swap_comm)
       
    99 done
       
   100 
       
   101 lemma subst_assoc: "subst (s1\<bullet>(s2\<bullet>s3))=subst ((s1\<bullet>s2)\<bullet>s3)"
       
   102 apply(rule ext)
       
   103 apply(induct_tac x)
       
   104 apply(auto)
       
   105 apply(simp add: subst_comp_expand)
       
   106 done
       
   107 
       
   108 lemma fresh_subst: "nabla1\<turnstile>a\<sharp>t\<Longrightarrow> nabla2\<Turnstile>(subst s) nabla1 \<longrightarrow> nabla2\<turnstile>a\<sharp>subst s t"
       
   109 apply(erule fresh.induct)
       
   110 apply(auto)
       
   111 --Susp
       
   112 apply(simp add: ext_subst_def subst_susp)
       
   113 apply(drule_tac x="(swapas (rev pi) a, X)" in bspec)
       
   114 apply(assumption)
       
   115 apply(simp)
       
   116 apply(rule fresh_swap_right[THEN mp])
       
   117 apply(assumption)
       
   118 done
       
   119 
       
   120 lemma equ_subst: "nabla1\<turnstile>t1\<approx>t2\<Longrightarrow> nabla2\<Turnstile> (subst s) nabla1 \<longrightarrow> nabla2\<turnstile>(subst s t1)\<approx>(subst s t2)"
       
   121 apply(erule equ.induct)
       
   122 apply(auto)
       
   123 apply(rule_tac "nabla1.1"="nabla" in fresh_subst[THEN mp])
       
   124 apply(assumption)+
       
   125 apply(simp add: subst_swap_comm)
       
   126 -- Susp
       
   127 apply(simp only: subst_susp)
       
   128 apply(rule equ_pi1_pi2_add[THEN mp])
       
   129 apply(simp only: ext_subst_def subst_susp)
       
   130 apply(force)
       
   131 done
       
   132 
       
   133 lemma unif_1: 
       
   134   "\<lbrakk>nabla\<turnstile>subst s (Susp pi X)\<approx>subst s t\<rbrakk>\<Longrightarrow> nabla\<Turnstile> subst (s\<bullet>[(X,swap (rev pi) t)])\<approx>subst s"
       
   135 apply(simp only: subst_equ_def)
       
   136 apply(case_tac "pi=[]")
       
   137 apply(simp add: subst_susp comp_def)
       
   138 apply(force intro: equ_sym equ_refl)
       
   139 apply(subgoal_tac "domn (subst (s\<bullet>[(X,swap (rev pi) t)]))=domn(subst s)\<union>{X}")--A
       
   140 apply(simp)
       
   141 apply(rule conjI)
       
   142 apply(simp add: subst_comp_expand)
       
   143 apply(simp add: subst_susp)
       
   144 apply(simp only: subst_swap_comm)
       
   145 apply(simp only: equ_pi_to_left)
       
   146 apply(rule equ_sym)
       
   147 apply(assumption)
       
   148 apply(rule ballI)
       
   149 apply(simp only: subst_comp_expand)
       
   150 apply(simp add: subst_susp)
       
   151 apply(force intro: equ_sym equ_refl simp add: subst_swap_comm equ_pi_to_left)
       
   152 --A
       
   153 apply(simp only: domn_def)
       
   154 apply(auto)
       
   155 apply(simp add: subst_comp_expand)
       
   156 apply(simp add: subst_susp subst_swap_comm)
       
   157 apply(simp only: subst_comp_expand)
       
   158 apply(simp add: subst_susp subst_swap_comm)
       
   159 apply(drule swap_empty[THEN mp])
       
   160 apply(simp)
       
   161 apply(simp only: subst_comp_expand)
       
   162 apply(simp only: subst_susp)
       
   163 apply(auto)
       
   164 apply(case_tac "x=X")
       
   165 apply(simp)
       
   166 apply(simp add: subst_swap_comm)
       
   167 apply(drule swap_empty[THEN mp])
       
   168 apply(simp)
       
   169 apply(simp add: subst_susp)
       
   170 done
       
   171 
       
   172 lemma subst_equ_a:
       
   173 "\<lbrakk>nabla\<Turnstile>(subst s1) \<approx> (subst s2)\<rbrakk>\<Longrightarrow> \<forall>t2. nabla\<turnstile>(subst s2 t1)\<approx>t2 \<longrightarrow> nabla\<turnstile>(subst s1 t1)\<approx>t2"
       
   174 apply(rule allI)
       
   175 apply(induct t1)
       
   176 --Abst.ab
       
   177 apply(rule impI)
       
   178 apply(simp)
       
   179 apply(ind_cases "nabla \<turnstile> Abst list (subst s1 trm) \<approx> t2")
       
   180 apply(best)
       
   181 --Abst.aa
       
   182 apply(force)
       
   183 --Susp
       
   184 apply(rule impI)
       
   185 apply(simp only: subst_equ_def)
       
   186 apply(case_tac "list2\<in>domn (subst s1) \<union> domn (subst s2)")
       
   187 apply(drule_tac x="list2" in bspec)
       
   188 apply(assumption)
       
   189 apply(simp)
       
   190 apply(subgoal_tac "nabla \<turnstile> subst s2 (Susp [] list2) \<approx> swap (rev list1) t2")--A
       
   191 apply(drule_tac "t1.0"="subst s1 (Susp [] list2)" and
       
   192                 "t2.0"="subst s2 (Susp [] list2)" and 
       
   193                 "t3.0"="swap (rev list1) t2" in equ_trans)
       
   194 apply(assumption)
       
   195 apply(simp only: equ_pi_to_right)
       
   196 apply(simp add: subst_swap_comm[THEN sym])
       
   197 --A
       
   198 apply(simp only: equ_pi_to_right)
       
   199 apply(simp add: subst_swap_comm[THEN sym])
       
   200 apply(simp)
       
   201 apply(erule conjE)
       
   202 apply(drule not_in_domn)+
       
   203 apply(subgoal_tac "subst s1 (Susp list1 list2)=swap list1 (subst s1 (Susp [] list2))")--B
       
   204 apply(subgoal_tac "subst s2 (Susp list1 list2)=swap list1 (subst s2 (Susp [] list2))")--C
       
   205 apply(simp)
       
   206 --BC
       
   207 apply(simp (no_asm) add: subst_swap_comm[THEN sym])
       
   208 apply(simp (no_asm) add: subst_swap_comm[THEN sym])
       
   209 --Unit
       
   210 apply(force)
       
   211 --Atom
       
   212 apply(force)
       
   213 --Paar
       
   214 apply(rule impI)
       
   215 apply(simp)
       
   216 apply(ind_cases "nabla \<turnstile> Paar (subst sigma1 trm1) (subst sigma1 trm2) \<approx> t2")
       
   217 apply(best)
       
   218 --Func
       
   219 apply(rule impI)
       
   220 apply(simp)
       
   221 apply(ind_cases "nabla \<turnstile> Func list (subst sigma1 trm) \<approx> t2")
       
   222 apply(best)
       
   223 done
       
   224 
       
   225 lemma unif_2a: "\<lbrakk>nabla\<Turnstile>subst s1\<approx>subst s2\<rbrakk>\<Longrightarrow> 
       
   226                 (nabla\<turnstile>subst s2 t1 \<approx> subst s2 t2)\<longrightarrow>(nabla\<turnstile>subst s1 t1 \<approx> subst s1 t2)"
       
   227 apply(rule impI)
       
   228 apply(frule_tac "t1.0"="t1" in subst_equ_a)
       
   229 apply(drule_tac x="subst s2 t2" in spec)
       
   230 apply(simp)
       
   231 apply(drule equ_sym)
       
   232 apply(drule equ_sym) 
       
   233 apply(frule_tac "t1.0"="t2" in subst_equ_a)
       
   234 apply(drule_tac x="subst s1 t1" in spec)
       
   235 apply(rule equ_sym)
       
   236 apply(simp)
       
   237 done
       
   238 
       
   239 
       
   240 lemma unif_2b: "\<lbrakk>nabla\<Turnstile>subst s1\<approx> subst s2\<rbrakk>\<Longrightarrow>nabla\<turnstile>a\<sharp> subst s2 t\<longrightarrow>nabla\<turnstile>a\<sharp>subst s1 t"
       
   241 apply(induct t)
       
   242 --Abst
       
   243 apply(simp)
       
   244 apply(rule impI)
       
   245 apply(case_tac "a=list")
       
   246 apply(force)
       
   247 apply(force dest!: fresh_abst_ab_elim)
       
   248 --Susp
       
   249 apply(rule impI)
       
   250 apply(subgoal_tac "subst s1 (Susp list1 list2)= swap list1 (subst s1 (Susp [] list2))")--A
       
   251 apply(subgoal_tac "subst s2 (Susp list1 list2)= swap list1 (subst s2 (Susp [] list2))")--B
       
   252 apply(simp add: subst_equ_def)
       
   253 apply(case_tac "list2\<in>domn (subst s1) \<union> domn (subst s2)")
       
   254 apply(drule_tac x="list2" in bspec)
       
   255 apply(assumption)
       
   256 apply(rule fresh_swap_right[THEN mp])
       
   257 apply(rule_tac "t1.1"=" subst s2 (Susp [] list2)" in l3_jud[THEN mp])
       
   258 apply(rule equ_sym)
       
   259 apply(simp)
       
   260 apply(rule fresh_swap_left[THEN mp])
       
   261 apply(simp)
       
   262 apply(simp)
       
   263 apply(erule conjE)
       
   264 apply(drule not_in_domn)+
       
   265 apply(simp add: subst_swap_comm)
       
   266 --AB
       
   267 apply(simp (no_asm) add: subst_swap_comm[THEN sym])
       
   268 apply(simp (no_asm) add: subst_swap_comm[THEN sym])
       
   269 --Unit
       
   270 apply(force)
       
   271 --Atom
       
   272 apply(force)
       
   273 --Paar
       
   274 apply(force dest!: fresh_paar_elim)
       
   275 --Func
       
   276 apply(force dest!:  fresh_func_elim)
       
   277 done
       
   278 
       
   279 lemma subst_equ_to_trm: "nabla\<Turnstile>subst s1 \<approx> subst s2\<Longrightarrow> nabla\<turnstile>subst s1 t\<approx>subst s2 t"
       
   280 apply(induct_tac t)
       
   281 apply(auto simp add: subst_equ_def)
       
   282 apply(case_tac "list2\<in>domn (subst s1) \<union> domn (subst s2)")
       
   283 apply(simp only: subst_susp)
       
   284 apply(simp only: equ_pi_to_left[THEN sym])
       
   285 apply(simp only: equ_involutive_left)
       
   286 apply(simp)
       
   287 apply(erule conjE)
       
   288 apply(drule not_in_domn)+
       
   289 apply(subgoal_tac "subst s1 (Susp list1 list2)=swap list1 (subst s1 (Susp [] list2))")--A
       
   290 apply(subgoal_tac "subst s2 (Susp list1 list2)=swap list1 (subst s2 (Susp [] list2))")--B
       
   291 apply(simp)
       
   292 apply(rule equ_refl)
       
   293 apply(simp (no_asm_use) add: subst_swap_comm[THEN sym])+
       
   294 done
       
   295 
       
   296 lemma subst_cancel_right: "\<lbrakk>nabla\<Turnstile>(subst s1)\<approx>(subst s2)\<rbrakk>\<Longrightarrow>nabla\<Turnstile>(subst (s1\<bullet>s))\<approx>(subst (s2\<bullet>s))" 
       
   297 apply(simp (no_asm) only: subst_equ_def)
       
   298 apply(rule ballI)
       
   299 apply(simp only: subst_comp_expand)
       
   300 apply(rule subst_equ_to_trm)
       
   301 apply(assumption)
       
   302 done
       
   303 
       
   304 lemma subst_trans: "\<lbrakk>nabla\<Turnstile>subst s1\<approx>subst s2; nabla\<Turnstile>subst s2\<approx>subst s3\<rbrakk>\<Longrightarrow>nabla\<Turnstile>subst s1\<approx>subst s3"
       
   305 apply(simp add: subst_equ_def)
       
   306 apply(auto)
       
   307 apply(case_tac "X \<in>domn (subst s2)")
       
   308 apply(drule_tac x="X" in bspec)
       
   309 apply(force)
       
   310 apply(drule_tac x="X" in bspec)
       
   311 apply(force)
       
   312 apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans)
       
   313 apply(assumption)+
       
   314 apply(case_tac "X \<in>domn (subst s3)")
       
   315 apply(drule_tac x="X" in bspec)
       
   316 apply(force)
       
   317 apply(drule_tac x="X" in bspec)
       
   318 apply(force)
       
   319 apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans)
       
   320 apply(assumption)+
       
   321 apply(drule not_in_domn)+
       
   322 apply(drule_tac x="X" in bspec)
       
   323 apply(force)
       
   324 apply(simp)
       
   325 apply(case_tac "X \<in>domn (subst s1)")
       
   326 apply(drule_tac x="X" in bspec)
       
   327 apply(force)
       
   328 apply(drule_tac x="X" in bspec)
       
   329 apply(force)
       
   330 apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans)
       
   331 apply(assumption)+
       
   332 apply(case_tac "X \<in>domn (subst s2)")
       
   333 apply(drule_tac x="X" in bspec)
       
   334 apply(force)
       
   335 apply(drule_tac x="X" in bspec)
       
   336 apply(force)
       
   337 apply(rule_tac "t2.0"="subst s2 (Susp [] X)" in equ_trans)
       
   338 apply(assumption)+
       
   339 apply(drule not_in_domn)+
       
   340 apply(simp)
       
   341 apply(rotate_tac 1)
       
   342 apply(drule_tac x="X" in bspec)
       
   343 apply(force)
       
   344 apply(simp)
       
   345 done
       
   346 
       
   347 (* if occurs holds, then one subterm is equal to (subst s (Susp pi X)) *)
       
   348 
       
   349 lemma occurs_sub_trm_equ: 
       
   350   "occurs X t1 \<longrightarrow> (\<exists>t2\<in>sub_trms (subst s t1). (\<exists>pi.(nabla\<turnstile>subst s (Susp pi1 X)\<approx>(swap pi t2))))" 
       
   351 apply(induct_tac t1)
       
   352 apply(auto)
       
   353 apply(simp only: subst_susp)
       
   354 apply(rule_tac x="swap list1 (look_up X s)" in bexI)
       
   355 apply(rule_tac x="pi1@(rev list1)" in exI)
       
   356 apply(simp add: swap_append)
       
   357 apply(simp add: equ_pi_to_left[THEN sym])
       
   358 apply(simp only: equ_involutive_left)
       
   359 apply(rule equ_refl)
       
   360 apply(rule t_sub_trms_t)
       
   361 done
       
   362 
       
   363 end