diff -r 32ec97f94a07 -r 2363eb91d9fd thys/rec_def.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys/rec_def.thy Wed Jan 23 20:18:40 2013 +0100 @@ -0,0 +1,87 @@ +theory rec_def +imports Main +begin + +section {* + Recursive functions +*} + +text {* + Datatype of recursive operators. +*} + +datatype recf = + -- {* The zero function, which always resturns @{text "0"} as result. *} + z | + -- {* The successor function, which increments its arguments. *} + s | + -- {* + The projection function, where @{text "id i j"} returns the @{text "j"}-th + argment out of the @{text "i"} arguments. + *} + id nat nat | + -- {* + The compostion operator, where "@{text "Cn n f [g1; g2; \ ;gm]"} + computes @{text "f (g1(x1, x2, \, xn), g2(x1, x2, \, xn), \ , + gm(x1, x2, \ , xn))"} for input argments @{text "x1, \, xn"}. + *} + Cn nat recf "recf list" | +-- {* + The primitive resursive operator, where @{text "Pr n f g"} computes: + @{text "Pr n f g (x1, x2, \, xn-1, 0) = f(x1, \, xn-1)"} + and @{text "Pr n f g (x1, x2, \, xn-1, k') = g(x1, x2, \, xn-1, k, + Pr n f g (x1, \, xn-1, k))"}. + *} + Pr nat recf recf | +-- {* + The minimization operator, where @{text "Mn n f (x1, x2, \ , xn)"} + computes the first i such that @{text "f (x1, \, xn, i) = 0"} and for all + @{text "j"}, @{text "f (x1, x2, \, xn, j) > 0"}. + *} + Mn nat recf + +text {* + The semantis of recursive operators is given by an inductively defined + relation as follows, where + @{text "rec_calc_rel R [x1, x2, \, xn] r"} means the computation of + @{text "R"} over input arguments @{text "[x1, x2, \, xn"} terminates + and gives rise to a result @{text "r"} +*} + +inductive rec_calc_rel :: "recf \ nat list \ nat \ bool" +where + calc_z: "rec_calc_rel z [n] 0" | + calc_s: "rec_calc_rel s [n] (Suc n)" | + calc_id: "\length args = i; j < i; args!j = r\ \ rec_calc_rel (id i j) args r" | + calc_cn: "\length args = n; + \ k < length gs. rec_calc_rel (gs ! k) args (rs ! k); + length rs = length gs; + rec_calc_rel f rs r\ + \ rec_calc_rel (Cn n f gs) args r" | + calc_pr_zero: + "\length args = n; + rec_calc_rel f args r0 \ + \ rec_calc_rel (Pr n f g) (args @ [0]) r0" | + calc_pr_ind: " + \ length args = n; + rec_calc_rel (Pr n f g) (args @ [k]) rk; + rec_calc_rel g (args @ [k] @ [rk]) rk'\ + \ rec_calc_rel (Pr n f g) (args @ [Suc k]) rk'" | + calc_mn: "\length args = n; + rec_calc_rel f (args@[r]) 0; + \ i < r. (\ ri. rec_calc_rel f (args@[i]) ri \ ri \ 0)\ + \ rec_calc_rel (Mn n f) args r" + +inductive_cases calc_pr_reverse: + "rec_calc_rel (Pr n f g) (lm) rSucy" + +inductive_cases calc_z_reverse: "rec_calc_rel z lm x" + +inductive_cases calc_s_reverse: "rec_calc_rel s lm x" + +inductive_cases calc_id_reverse: "rec_calc_rel (id m n) lm x" + +inductive_cases calc_cn_reverse: "rec_calc_rel (Cn n f gs) lm x" + +inductive_cases calc_mn_reverse:"rec_calc_rel (Mn n f) lm x" +end \ No newline at end of file