Unification/Terms.thy
changeset 107 5c816239deaa
equal deleted inserted replaced
106:ed54ec416bb3 107:5c816239deaa
       
     1 
       
     2 theory Terms = Main + Swap:
       
     3 
       
     4 (* terms *)
       
     5 
       
     6 datatype trm = Abst  string  trm   
       
     7              | Susp  "(string \<times> string)list" string
       
     8              | Unit                                   
       
     9              | Atom  string                           
       
    10              | Paar  trm trm                          
       
    11              | Func string trm                       
       
    12 
       
    13 (* swapping operation on terms *)
       
    14 
       
    15 consts swap  :: "(string \<times> string)list \<Rightarrow> trm \<Rightarrow> trm"
       
    16 primrec
       
    17 "swap  pi (Unit)        = Unit"
       
    18 "swap  pi (Atom a)      = Atom (swapas pi a)"
       
    19 "swap  pi (Susp pi' X)  = Susp (pi@pi') X"
       
    20 "swap  pi (Abst a t)    = Abst (swapas pi a) (swap pi t)"
       
    21 "swap  pi (Paar t1 t2)  = Paar (swap pi t1) (swap pi t2)"
       
    22 "swap  pi (Func F t)    = Func F (swap pi t)"
       
    23 
       
    24 lemma [simp]: "swap [] t = t"
       
    25 apply(induct_tac t)
       
    26 apply(auto)
       
    27 done
       
    28 
       
    29 lemma swap_append[rule_format]: "\<forall>pi1 pi2. swap (pi1@pi2) t = swap pi1 (swap pi2 t)"
       
    30 apply(induct_tac t, auto)
       
    31 apply(induct_tac pi1,auto)
       
    32 apply(induct_tac pi1,auto)
       
    33 done
       
    34 
       
    35 lemma swap_empty: "swap pi t=Susp [] X\<longrightarrow> pi=[]"
       
    36 apply(induct_tac t)
       
    37 apply(auto)
       
    38 done
       
    39 
       
    40 (* depyh of terms *)
       
    41 
       
    42 consts depth :: "trm \<Rightarrow> nat"
       
    43 primrec
       
    44 "depth (Unit)      = 0"
       
    45 "depth (Atom a)    = 0"
       
    46 "depth (Susp pi X) = 0"
       
    47 "depth (Abst a t)  = 1 + depth t"
       
    48 "depth (Func F t)  = 1 + depth t"
       
    49 "depth (Paar t t') = 1 + (max (depth t) (depth t'))" 
       
    50 
       
    51 lemma 
       
    52   Suc_max_left:  "Suc(max x y) = z \<longrightarrow> x < z" and
       
    53   Suc_max_right: "Suc(max x y) = z \<longrightarrow> y < z"
       
    54 apply(auto)
       
    55 apply(simp_all only: max_Suc_Suc[THEN sym] less_max_iff_disj)
       
    56 apply(simp_all)
       
    57 done
       
    58 
       
    59 lemma [simp]: "depth (swap pi t) = depth t" 
       
    60 apply(induct_tac t,auto)
       
    61 done
       
    62 
       
    63 (* occurs predicate and variables in terms *)
       
    64 
       
    65 consts occurs :: "string \<Rightarrow> trm \<Rightarrow> bool"
       
    66 primrec 
       
    67 "occurs X (Unit)       = False"
       
    68 "occurs X (Atom a)     = False"
       
    69 "occurs X (Susp pi Y)  = (if X=Y then True else False)"
       
    70 "occurs X (Abst a t)   = occurs X t"
       
    71 "occurs X (Paar t1 t2) = (if (occurs X t1) then True else (occurs X t2))"
       
    72 "occurs X (Func F t)   = occurs X t"
       
    73 
       
    74 consts vars_trm :: "trm \<Rightarrow> string set"
       
    75 primrec
       
    76 "vars_trm (Unit)       = {}"
       
    77 "vars_trm (Atom a)     = {}"
       
    78 "vars_trm (Susp pi X)  = {X}"
       
    79 "vars_trm (Paar t1 t2) = (vars_trm t1)\<union>(vars_trm t2)"
       
    80 "vars_trm (Abst a t)   = vars_trm t"
       
    81 "vars_trm (Func F t)   = vars_trm t"
       
    82 
       
    83 
       
    84 (* subterms and proper subterms *)
       
    85 
       
    86 consts sub_trms :: "trm \<Rightarrow> trm set"
       
    87 primrec 
       
    88 "sub_trms (Unit)       = {Unit}"
       
    89 "sub_trms (Atom a)     = {Atom a}"
       
    90 "sub_trms (Susp pi Y)  = {Susp pi Y}"
       
    91 "sub_trms (Abst a t)   = {Abst a t}\<union>sub_trms t"
       
    92 "sub_trms (Paar t1 t2) = {Paar t1 t2}\<union>sub_trms t1 \<union>sub_trms t2"
       
    93 "sub_trms (Func F t)   = {Func F t}\<union>sub_trms t"
       
    94 
       
    95 consts psub_trms :: "trm \<Rightarrow> trm set"
       
    96 primrec
       
    97 "psub_trms (Unit)       = {}"
       
    98 "psub_trms (Atom a)     = {}"
       
    99 "psub_trms (Susp pi X)  = {}"
       
   100 "psub_trms (Paar t1 t2) = sub_trms t1 \<union> sub_trms t2"
       
   101 "psub_trms (Abst a t)   = sub_trms t"
       
   102 "psub_trms (Func F t)   = sub_trms t"
       
   103 
       
   104 lemma psub_sub_trms[rule_format]: 
       
   105   "\<forall>t1 . t1 \<in> psub_trms t2 \<longrightarrow>  t1 \<in> sub_trms t2"
       
   106 apply(induct t2)
       
   107 apply(auto)
       
   108 done
       
   109 
       
   110 lemma t_sub_trms_t: 
       
   111   "t\<in> sub_trms t"
       
   112 apply(induct t)
       
   113 apply(auto)
       
   114 done
       
   115 
       
   116 lemma abst_psub_trms[rule_format]: 
       
   117   "\<forall> t1. Abst a t1 \<in> sub_trms t2 \<longrightarrow> t1 \<in> psub_trms t2"
       
   118 apply(induct_tac t2)
       
   119 apply(auto)
       
   120 apply(simp only: t_sub_trms_t)
       
   121 apply(best dest!: psub_sub_trms)+
       
   122 done
       
   123 
       
   124 lemma func_psub_trms[rule_format]: 
       
   125   "\<forall> t1. Func F t1 \<in> sub_trms t2 \<longrightarrow> t1 \<in> psub_trms t2"
       
   126 apply(induct_tac t2)
       
   127 apply(auto)
       
   128 apply(best dest!: psub_sub_trms)+
       
   129 apply(simp add: t_sub_trms_t)
       
   130 apply(best dest!: psub_sub_trms)
       
   131 done
       
   132 
       
   133 lemma paar_psub_trms[rule_format]: 
       
   134   "\<forall> t1. Paar t1 t2 \<in> sub_trms t3\<longrightarrow>(t1 \<in> psub_trms t3 \<and> t2 \<in> psub_trms t3)"
       
   135 apply(induct_tac t3)
       
   136 apply(auto)
       
   137 apply(best dest!: psub_sub_trms)+
       
   138 apply(simp add: t_sub_trms_t)+
       
   139 apply(best dest!: psub_sub_trms)+
       
   140 done
       
   141 
       
   142 end