Unification/Termination.thy
changeset 107 5c816239deaa
equal deleted inserted replaced
106:ed54ec416bb3 107:5c816239deaa
       
     1 
       
     2 
       
     3 theory Termination = Main + Terms + Fresh + Equ + Substs + Mgu:
       
     4 
       
     5 (* set of variables *)
       
     6 
       
     7 consts 
       
     8   vars_trm    :: "trm \<Rightarrow> string set"
       
     9   vars_eprobs :: "eprobs \<Rightarrow> (string set)"
       
    10 primrec
       
    11   "vars_trm (Unit)       = {}"
       
    12   "vars_trm (Atom a)     = {}"
       
    13   "vars_trm (Susp pi X)  = {X}"
       
    14   "vars_trm (Paar t1 t2) = (vars_trm t1)\<union>(vars_trm t2)"
       
    15   "vars_trm (Abst a t)   = vars_trm t"
       
    16   "vars_trm (Func F t)   = vars_trm t"
       
    17 primrec
       
    18   "vars_eprobs [] = {}"
       
    19   "vars_eprobs (x#xs) = (vars_trm (snd x))\<union>(vars_trm (fst x))\<union>(vars_eprobs xs)"
       
    20 
       
    21 lemma[simp]: "vars_trm (swap pi t) = vars_trm t"
       
    22 apply(induct_tac t)
       
    23 apply(auto)
       
    24 done
       
    25 
       
    26 consts 
       
    27   size_trm    :: "trm \<Rightarrow> nat"
       
    28   size_fprobs :: "fprobs \<Rightarrow> nat"
       
    29   size_eprobs :: "eprobs \<Rightarrow> nat"
       
    30   size_probs  :: "probs \<Rightarrow> nat"
       
    31 primrec
       
    32   "size_trm (Unit)      = 1"
       
    33   "size_trm (Atom a)    = 1"
       
    34   "size_trm (Susp pi X) = 1"
       
    35   "size_trm (Abst a t)  = 1 + size_trm t"
       
    36   "size_trm (Func F t)  = 1 + size_trm t"
       
    37   "size_trm (Paar t t') = 1 + (size_trm t) + (size_trm t')"
       
    38 primrec
       
    39   "size_fprobs [] = 0"
       
    40   "size_fprobs (x#xs) = (size_trm (snd x))+(size_fprobs xs)"
       
    41 primrec
       
    42   "size_eprobs [] = 0"
       
    43   "size_eprobs (x#xs) = (size_trm (fst x))+(size_trm (snd x))+(size_eprobs xs)"
       
    44 
       
    45 lemma[simp]: "size_trm (swap pi t) = size_trm t"
       
    46 apply(induct_tac t)
       
    47 apply(auto)
       
    48 done
       
    49 
       
    50 syntax        
       
    51   "_measure_relation" :: "(nat\<times>nat\<times>nat) \<Rightarrow> (nat\<times>nat\<times>nat) \<Rightarrow> bool" ("_ \<lless> _" [80,80] 80)
       
    52 translations
       
    53   "n1 \<lless> n2"  \<rightleftharpoons> "(n1,n2) \<in> (less_than<*lex*>less_than<*lex*>less_than)"
       
    54 
       
    55 consts 
       
    56   rank :: "probs \<Rightarrow> (nat\<times>nat\<times>nat)"
       
    57 primrec
       
    58   "rank (eprobs,fprobs) = (card (vars_eprobs eprobs),size_eprobs eprobs,size_fprobs fprobs)" 
       
    59 
       
    60 lemma[simp]: "finite (vars_trm t)"
       
    61 apply(induct t)
       
    62 apply(auto)
       
    63 done
       
    64 
       
    65 lemma[simp]: "finite (vars_eprobs P)"
       
    66 apply(induct_tac P)
       
    67 apply(auto)
       
    68 done
       
    69 
       
    70 lemma union_comm: "A\<union>(B\<union>C)=(A\<union>B)\<union>C"
       
    71 apply(auto)
       
    72 done
       
    73 
       
    74 lemma card_union: "\<lbrakk>finite A; finite B\<rbrakk>\<Longrightarrow>(card B < card (A\<union>B)) \<or> (card B = card (A\<union>B))"
       
    75 apply(auto)
       
    76 apply(rule psubset_card_mono)
       
    77 apply(auto)
       
    78 done
       
    79 
       
    80 lemma card_insert: "\<lbrakk>finite B\<rbrakk>\<Longrightarrow>(card B < card (insert X B)) \<or> (card B = card (insert X B))"
       
    81 apply(auto)
       
    82 apply(rule psubset_card_mono)
       
    83 apply(auto)
       
    84 done
       
    85 
       
    86 lemma subseteq_card: "\<lbrakk>A\<subseteq>B; finite B\<rbrakk>\<Longrightarrow>(card A \<le> card B)"
       
    87 apply(drule_tac A="A" in card_mono)
       
    88 apply(auto simp add: le_eq_less_or_eq)
       
    89 done
       
    90 
       
    91 lemma not_occurs_trm: "\<not>occurs X t \<longrightarrow> X\<notin> vars_trm t"
       
    92 apply(induct_tac t)
       
    93 apply(auto)
       
    94 done
       
    95 
       
    96 lemma not_occurs_subst: "\<not>occurs X t1\<longrightarrow> X\<notin> vars_trm (subst [(X,swap pi2 t1)] t2)" 
       
    97 apply(induct_tac t2)
       
    98 apply(auto simp add: subst_susp not_occurs_trm)
       
    99 done
       
   100 
       
   101 lemma not_occurs_list: "\<not>occurs X t \<longrightarrow>
       
   102   X\<notin> vars_eprobs (apply_subst_eprobs [(X, swap pi t)] xs)"
       
   103 apply(induct_tac xs)
       
   104 apply(simp)
       
   105 apply(case_tac a)
       
   106 apply(auto simp add: not_occurs_subst)
       
   107 done
       
   108 
       
   109 lemma vars_equ: "\<not>occurs X t1 \<and> occurs X t2\<longrightarrow> 
       
   110   vars_trm (subst [(X, swap pi t1)] t2)=(vars_trm t1\<union>vars_trm t2)-{X}"
       
   111 apply(induct_tac t2)
       
   112 apply(force)
       
   113 apply(case_tac "X=list2")
       
   114 apply(simp add: subst_susp not_occurs_trm)
       
   115 apply(simp)
       
   116 apply(simp)
       
   117 apply(simp)
       
   118 apply(simp)
       
   119 apply(rule conjI)
       
   120 apply(case_tac "occurs X trm2")
       
   121 apply(force)
       
   122 apply(force dest: not_occurs_trm[THEN mp] simp add: subst_not_occurs)
       
   123 apply(force dest: not_occurs_trm[THEN mp] simp add: subst_not_occurs)
       
   124 apply(force)
       
   125 done
       
   126 
       
   127 lemma vars_subseteq: "\<not>occurs X t \<longrightarrow>
       
   128   vars_eprobs (apply_subst_eprobs [(X, swap pi t)] xs) \<subseteq> (vars_trm t \<union> vars_eprobs xs)"
       
   129 apply(induct_tac xs)
       
   130 apply(simp)
       
   131 apply(rule impI)
       
   132 apply(simp)
       
   133 apply(case_tac "occurs X (fst a)")
       
   134 apply(case_tac "occurs X (snd a)")
       
   135 apply(simp add: vars_equ[THEN mp])
       
   136 apply(force)
       
   137 apply(simp add: subst_not_occurs)
       
   138 apply(force simp add: vars_equ)
       
   139 apply(case_tac "occurs X (snd a)")
       
   140 apply(simp add: vars_equ[THEN mp])
       
   141 apply(force simp add: subst_not_occurs)
       
   142 apply(force simp add: subst_not_occurs)
       
   143 done
       
   144 
       
   145 lemma vars_decrease: 
       
   146   "\<not>occurs X t\<longrightarrow> card (vars_eprobs (apply_subst_eprobs [(X, swap pi t)] xs))
       
   147                  <card (insert X (vars_trm t \<union> vars_eprobs xs))"
       
   148 apply(rule impI)
       
   149 apply(case_tac "X\<in>(vars_trm t \<union> vars_eprobs xs)")
       
   150 (* first case *)
       
   151 apply(subgoal_tac "insert X (vars_trm t \<union> vars_eprobs xs) = (vars_trm t \<union> vars_eprobs xs)") (*A*)
       
   152 apply(simp)
       
   153 apply(frule_tac pi1="pi" and xs1="xs" in vars_subseteq[THEN mp])
       
   154 apply(frule_tac pi1="pi" and xs1="xs" in not_occurs_list[THEN mp])
       
   155 apply(subgoal_tac "vars_eprobs (apply_subst_eprobs [(X, swap pi t)] xs)
       
   156                    \<subset>  vars_trm t \<union> vars_eprobs xs") (* B *)
       
   157 apply(simp add: psubset_card_mono)
       
   158 (* B *)
       
   159 apply(force)
       
   160 (* A *)
       
   161 apply(force)
       
   162 (* second case *)
       
   163 apply(subgoal_tac "finite (vars_trm t \<union> vars_eprobs xs)")
       
   164 apply(drule_tac x="X" in card_insert_disjoint)
       
   165 apply(simp)
       
   166 apply(simp)
       
   167 apply(frule_tac pi1="pi" and xs1="xs" in vars_subseteq[THEN mp])
       
   168 apply(auto dest: subseteq_card)
       
   169 done
       
   170 
       
   171 lemma rank_cred: "\<lbrakk>P1\<turnstile>(nabla)\<rightarrow>P2\<rbrakk>\<Longrightarrow>(rank P2) \<lless> (rank P1)"
       
   172 apply(ind_cases "P1 \<turnstile> nabla \<rightarrow> P2")
       
   173 apply(simp_all add: lex_prod_def)
       
   174 done
       
   175 
       
   176 lemma rank_sred: "\<lbrakk>P1\<turnstile>(s)\<leadsto>P2\<rbrakk>\<Longrightarrow>(rank P2) \<lless> (rank P1)"
       
   177 apply(ind_cases "P1 \<turnstile> s \<leadsto> P2")
       
   178 apply(simp_all add: lex_prod_def union_comm)
       
   179 apply(subgoal_tac 
       
   180   "vars_trm s1\<union>vars_trm t1\<union>vars_trm s2\<union>vars_trm t2\<union>vars_eprobs xs = 
       
   181    vars_trm s1\<union>vars_trm s2\<union>vars_trm t1\<union>vars_trm t2\<union>vars_eprobs xs") (*A*)
       
   182 apply(simp)
       
   183 (* A *)
       
   184 apply(force)
       
   185 (* Susp-Susp case *)
       
   186 apply(simp add: card_insert)
       
   187 (* variable elimination cases *)
       
   188 apply(simp_all add: apply_subst_def vars_decrease)
       
   189 done
       
   190 
       
   191 
       
   192 lemma rank_trans: "\<lbrakk>rank P1 \<lless> rank P2; rank P2 \<lless> rank P3\<rbrakk>\<Longrightarrow>rank P1 \<lless> rank P3"
       
   193 apply(simp add: lex_prod_def)
       
   194 apply(auto)
       
   195 done
       
   196 
       
   197 (* all reduction are well-founded under \<lless> *)
       
   198 
       
   199 lemma rank_red_plus: "\<lbrakk>P1\<Turnstile>(s,nabla)\<Rightarrow>P2\<rbrakk>\<Longrightarrow>(rank P2) \<lless> (rank P1)"
       
   200 apply(erule red_plus.induct)
       
   201 apply(auto dest: rank_sred rank_cred rank_trans)
       
   202 done
       
   203 
       
   204 end
       
   205