utm/rec_def.thy
author zhang
Sat, 29 Sep 2012 12:38:12 +0000
changeset 370 1ce04eb1c8ad
permissions -rw-r--r--
Initial upload of the formal construction of Universal Turing Machine.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
370
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     1
theory rec_def
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     2
imports Main
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     3
begin
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     5
section {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     6
  Recursive functions
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     7
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     8
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     9
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    10
  Datatype of recursive operators.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    11
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    12
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    13
datatype recf = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    14
 -- {* The zero function, which always resturns @{text "0"} as result. *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    15
  z | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    16
 -- {* The successor function, which increments its arguments. *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    17
  s | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    18
 -- {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    19
  The projection function, where @{text "id i j"} returns the @{text "j"}-th
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    20
  argment out of the @{text "i"} arguments.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    21
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    22
  id nat nat | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    23
 -- {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    24
  The compostion operator, where "@{text "Cn n f [g1; g2; \<dots> ;gm]"} 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    25
  computes @{text "f (g1(x1, x2, \<dots>, xn), g2(x1, x2, \<dots>, xn), \<dots> , 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    26
  gm(x1, x2, \<dots> , xn))"} for input argments @{text "x1, \<dots>, xn"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    27
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    28
  Cn nat recf "recf list" | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    29
-- {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    30
  The primitive resursive operator, where @{text "Pr n f g"} computes:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    31
  @{text "Pr n f g (x1, x2, \<dots>, xn-1, 0) = f(x1, \<dots>, xn-1)"} 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    32
  and @{text "Pr n f g (x1, x2, \<dots>, xn-1, k') = g(x1, x2, \<dots>, xn-1, k, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    33
                                                  Pr n f g (x1, \<dots>, xn-1, k))"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    34
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    35
  Pr nat recf recf | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    36
-- {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    37
  The minimization operator, where @{text "Mn n f (x1, x2, \<dots> , xn)"} 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    38
  computes the first i such that @{text "f (x1, \<dots>, xn, i) = 0"} and for all
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    39
  @{text "j"}, @{text "f (x1, x2, \<dots>, xn, j) > 0"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    40
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    41
  Mn nat recf 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    42
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    43
text {* 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    44
  The semantis of recursive operators is given by an inductively defined
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    45
  relation as follows, where  
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    46
  @{text "rec_calc_rel R [x1, x2, \<dots>, xn] r"} means the computation of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    47
  @{text "R"} over input arguments @{text "[x1, x2, \<dots>, xn"} terminates
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    48
  and gives rise to a result @{text "r"}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    49
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    50
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    51
inductive rec_calc_rel :: "recf \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    52
where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    53
  calc_z: "rec_calc_rel z [n] 0" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    54
  calc_s: "rec_calc_rel s [n] (Suc n)" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    55
  calc_id: "\<lbrakk>length args = i; j < i; args!j = r\<rbrakk> \<Longrightarrow> rec_calc_rel (id i j) args r" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    56
  calc_cn: "\<lbrakk>length args = n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    57
             \<forall> k < length gs. rec_calc_rel (gs ! k) args (rs ! k);
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    58
             length rs = length gs; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    59
             rec_calc_rel f rs r\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    60
            \<Longrightarrow> rec_calc_rel (Cn n f gs) args r" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    61
  calc_pr_zero: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    62
           "\<lbrakk>length args = n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    63
             rec_calc_rel f args r0 \<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    64
            \<Longrightarrow> rec_calc_rel (Pr n f g) (args @ [0]) r0" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    65
  calc_pr_ind: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    66
           \<lbrakk> length args = n;
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    67
             rec_calc_rel (Pr n f g) (args @ [k]) rk; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    68
             rec_calc_rel g (args @ [k] @ [rk]) rk'\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    69
            \<Longrightarrow> rec_calc_rel (Pr n f g) (args @ [Suc k]) rk'"  |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    70
  calc_mn: "\<lbrakk>length args = n; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    71
             rec_calc_rel f (args@[r]) 0; 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    72
             \<forall> i < r. (\<exists> ri. rec_calc_rel f (args@[i]) ri \<and> ri \<noteq> 0)\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    73
            \<Longrightarrow> rec_calc_rel (Mn n f) args r" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    74
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    75
inductive_cases calc_pr_reverse:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    76
              "rec_calc_rel (Pr n f g) (lm) rSucy"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    77
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    78
inductive_cases calc_z_reverse: "rec_calc_rel z lm x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    79
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    80
inductive_cases calc_s_reverse: "rec_calc_rel s lm x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    81
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    82
inductive_cases calc_id_reverse: "rec_calc_rel (id m n) lm x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    83
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    84
inductive_cases calc_cn_reverse: "rec_calc_rel (Cn n f gs) lm x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    85
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    86
inductive_cases calc_mn_reverse:"rec_calc_rel (Mn n f) lm x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    87
end