src/HOL/Library/SCT_Examples.thy
changeset 22359 94a794672c8b
child 22371 c9f5895972b0
equal deleted inserted replaced
22358:4d8a9e3a29f8 22359:94a794672c8b
       
     1 theory SCT_Examples
       
     2 imports Size_Change_Termination
       
     3 begin
       
     4 
       
     5 
       
     6 
       
     7 function f :: "nat \<Rightarrow> nat \<Rightarrow> nat"
       
     8 where
       
     9   "f n 0 = n"
       
    10 | "f 0 (Suc m) = f (Suc m) m"
       
    11 | "f (Suc n) (Suc m) = f m n"
       
    12 by pat_completeness auto
       
    13 
       
    14 
       
    15 termination
       
    16   unfolding f_rel_def lfp_const
       
    17   apply (rule SCT_on_relations)
       
    18   apply (tactic "SCT.abs_rel_tac") (* Build call descriptors *)
       
    19   apply (rule ext, rule ext, simp) (* Show that they are correct *)
       
    20   apply (tactic "SCT.mk_call_graph") (* Build control graph *)
       
    21   apply (rule LJA_apply)                 (* Apply main theorem *)
       
    22   apply (simp add:finite_acg_ins finite_acg_empty) (* show finiteness *)
       
    23   apply (rule SCT'_exec)
       
    24   by eval (* Evaluate to true *)
       
    25 
       
    26 
       
    27 function p :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
    28 where
       
    29   "p m n r = (if r>0 then p m (r - 1) n else
       
    30               if n>0 then p r (n - 1) m 
       
    31                      else m)"
       
    32 by pat_completeness auto
       
    33 
       
    34 termination
       
    35   unfolding p_rel_def lfp_const
       
    36   apply (rule SCT_on_relations)
       
    37   apply (tactic "SCT.abs_rel_tac") 
       
    38   apply (rule ext, rule ext, simp) 
       
    39   apply (tactic "SCT.mk_call_graph")
       
    40   apply (rule LJA_apply)            
       
    41   apply (simp add:finite_acg_ins finite_acg_empty) 
       
    42   apply (rule SCT'_exec)
       
    43   by eval
       
    44 
       
    45 function foo :: "bool \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
    46 where
       
    47   "foo True (Suc n) m = foo True n (Suc m)"
       
    48   "foo True 0 m = foo False 0 m"
       
    49   "foo False n (Suc m) = foo False (Suc n) m"
       
    50   "foo False n 0 = n"
       
    51 by pat_completeness auto
       
    52 
       
    53 termination
       
    54   unfolding foo_rel_def lfp_const
       
    55   apply (rule SCT_on_relations)
       
    56   apply (tactic "SCT.abs_rel_tac") 
       
    57   apply (rule ext, rule ext, simp) 
       
    58   apply (tactic "SCT.mk_call_graph")
       
    59   apply (rule LJA_apply)            
       
    60   apply (simp add:finite_acg_ins finite_acg_empty) 
       
    61   apply (rule SCT'_exec)
       
    62   by eval
       
    63 
       
    64 
       
    65 function (sequential) 
       
    66   bar :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
    67 where
       
    68   "bar 0 (Suc n) m = bar m m m"
       
    69 | "bar k n m = 0"
       
    70 by pat_completeness auto
       
    71 
       
    72 termination
       
    73   unfolding bar_rel_def lfp_const
       
    74   apply (rule SCT_on_relations)
       
    75   apply (tactic "SCT.abs_rel_tac") 
       
    76   apply (rule ext, rule ext, simp) 
       
    77   apply (tactic "SCT.mk_call_graph")
       
    78   apply (rule LJA_apply)            
       
    79   apply (simp add:finite_acg_ins finite_acg_empty) 
       
    80   by (rule SCT'_empty)
       
    81 
       
    82 
       
    83 end