src/Tools/isac/Knowledge/InsSort.thy
author Walther Neuper <neuper@ist.tugraz.at>
Fri, 10 Sep 2010 11:58:46 +0200
branchisac-update-Isa09-2
changeset 38002 10a171ce75d5
parent 37967 bd4f7a35e892
child 40836 69364e021751
permissions -rw-r--r--
intermediate in Knowledge/Isac.thy

adapted all code to Theory.axioms_of, which returns terms instead of thms
Isac.thy not finished
     1 (* 6.8.02 change to Isabelle2002 caused error -- thy excluded !
     2 
     3 Proving equations for primrec function(s) "InsSort.foldr" ...
     4 GC #1.17.30.54.345.21479:   (10 ms)
     5 *** Definition of InsSort.ins :: "['a::ord list, 'a::ord] => 'a::ord list"
     6 *** imposes additional sort constraints on the declared type of the constant
     7 *** The error(s) above occurred in definition "InsSort.ins.ins_list_def (@@@)"
     8 *)
     9 
    10 (* insertion sort, would need lists different from script-lists WN.11.00
    11 WN.7.5.03: -"- started with someList :: 'a list => unl, fun dest_list
    12 WN.8.5.03: error (@@@) remained with outcommenting foldr ?!?
    13 
    14 *)
    15 
    16 theory InsSort imports "../ProgLang/Script" begin
    17 
    18 consts
    19 
    20 (*foldr      :: [['a,'b] => 'a, 'b list, 'a] => 'a
    21 WN.8.5.03: already defined in Isabelle2002 (instantiated by Typefix):
    22      "[[real, real] => real, real list, real] => real") : term
    23 
    24  val t = str2term "foldr";
    25 val t =
    26   Const
    27     ("List.foldr",
    28      "[[RealDef.real, RealDef.real] => RealDef.real, RealDef.real List.list,
    29       RealDef.real] => RealDef.real") : term
    30  *)
    31   ins        :: ['a list,'a] => 'a list
    32   sort       :: 'a list => 'a list
    33 
    34 (*descriptions, script-id*)
    35   unsorted   :: 'a list => unl
    36   sorted     :: 'a list => unl
    37 
    38 (*subproblem and script-name*)
    39   Ins'_sort  :: "['a list,  
    40 		    'a list] => 'a list"
    41                ("((Script Ins'_sort (_ =))//  
    42 		    (_))" 9)
    43   Sort       :: "['a list,  
    44 		    'a list] => 'a list"
    45                ("((Script Sort (_ =))//  
    46 		    (_))" 9)
    47 
    48 (*primrec
    49   foldr_base "foldr f [] a = a"
    50   foldr_rec  "foldr f (x#xs) a = foldr f xs (f a x)"
    51 *)
    52 
    53 rules
    54 
    55 (*primrec .. outcommented analoguous to ListC.thy*)
    56   ins_base   "ins [] a = [a]"
    57   ins_rec    "ins (x#xs) a = (if x < a then x#(ins xs a) else a#(x#xs))" 
    58  
    59   sort_def   "sort ls = foldr ins ls []"
    60 
    61 ML {*
    62 (** problem type **)
    63 
    64 store_pbt
    65  (prep_pbt (theory "InsSort")
    66  (["functional"]:pblID,
    67   [("#Given" ,["unsorted u_"]),
    68    ("#Find"  ,["sorted s_"])
    69   ],
    70   []));
    71 
    72 store_pbt
    73  (prep_pbt (theory "InsSort")
    74  (["inssort","functional"]:pblID,
    75   [("#Given" ,["unsorted u_"]),
    76    ("#Find"  ,["sorted s_"])
    77   ],
    78   []));
    79 
    80 (** method, 
    81     todo: implementation needs extra object-level lists **)
    82 
    83 store_met
    84  (prep_met (theory "Diff")
    85  (["InsSort"],
    86    [],
    87    {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = e_rls, prls=e_rls,
    88     crls = Atools_rls, nrls=norm_Rational
    89     (*, asm_rls=[],asm_thm=[]*)}, "empty_script"));
    90 store_met
    91  (prep_met (theory "InsSort") (*test-version for [#1,#3,#2] only: see *.sml*)
    92  (["InsSort""sort"]:metID,
    93    [("#Given" ,["unsorted u_"]),
    94     ("#Find"  ,["sorted s_"])
    95     ],
    96    {rew_ord'="tless_true",rls'=eval_rls,calc = [], srls = e_rls, prls=e_rls,
    97     crls = eval_rls, nrls=norm_Rational(*,asm_rls=[],asm_thm=[]*)},
    98    "Script Sort (u_::'a list) = (Rewrite_Set ins_sort False) u_"
    99   ));
   100 
   101 ruleset' := overwritelthy @{theory} (!ruleset',
   102 			[(*("ins_sort",ins_sort) overwrites a Isa fun!!*)
   103 			 ]:(string * rls) list);
   104 *}
   105 
   106 end