src/HOL/Spec_Check/Examples.thy
changeset 53173 b881bee69d3a
parent 53159 a5805fe4e91c
parent 53172 31e24d6ff1ea
child 53174 71a2702da5e0
equal deleted inserted replaced
53159:a5805fe4e91c 53173:b881bee69d3a
     1 theory Examples
       
     2 imports Spec_Check
       
     3 begin
       
     4 
       
     5 section {* List examples *}
       
     6 
       
     7 ML_command {*
       
     8 check_property "ALL xs. rev xs = xs";
       
     9 *}
       
    10 
       
    11 ML_command {*
       
    12 check_property "ALL xs. rev (rev xs) = xs";
       
    13 *}
       
    14 
       
    15 
       
    16 section {* AList Specification *}
       
    17 
       
    18 ML_command {*
       
    19 (* map_entry applies the function to the element *)
       
    20 check_property "ALL k f xs. AList.lookup (op =) (AList.map_entry (op =) k f xs) k = Option.map f (AList.lookup (op =) xs k)";
       
    21 *}
       
    22 
       
    23 ML_command {*
       
    24 (* update always results in an entry *)
       
    25 check_property "ALL k v xs. AList.defined (op =) (AList.update (op =) (k, v) xs) k";
       
    26 *}
       
    27 
       
    28 ML_command {*
       
    29 (* update always writes the value *)
       
    30 check_property "ALL k v xs. AList.lookup (op =) (AList.update (op =) (k, v) xs) k = SOME v";
       
    31 *}
       
    32 
       
    33 ML_command {*
       
    34 (* default always results in an entry *)
       
    35 check_property "ALL k v xs. AList.defined (op =) (AList.default (op =) (k, v) xs) k";
       
    36 *}
       
    37 
       
    38 ML_command {*
       
    39 (* delete always removes the entry *)
       
    40 check_property "ALL k xs. not (AList.defined (op =) (AList.delete (op =) k xs) k)";
       
    41 *}
       
    42 
       
    43 ML_command {*
       
    44 (* default writes the entry iff it didn't exist *)
       
    45 check_property "ALL k v xs. (AList.lookup (op =) (AList.default (op =) (k, v) xs) k = (if AList.defined (op =) xs k then AList.lookup (op =) xs k else SOME v))";
       
    46 *}
       
    47 
       
    48 section {* Examples on Types and Terms *}
       
    49 
       
    50 ML_command {*
       
    51 check_property "ALL f g t. map_types (g o f) t = (map_types f o map_types g) t";
       
    52 *}
       
    53 
       
    54 ML_command {*
       
    55 check_property "ALL f g t. map_types (f o g) t = (map_types f o map_types g) t";
       
    56 *}
       
    57 
       
    58 
       
    59 text {* One would think this holds: *}
       
    60 
       
    61 ML_command {*
       
    62 check_property "ALL t ts. strip_comb (list_comb (t, ts)) = (t, ts)"
       
    63 *}
       
    64 
       
    65 text {* But it only holds with this precondition: *}
       
    66 
       
    67 ML_command {*
       
    68 check_property "ALL t ts. case t of _ $ _ => true | _ => strip_comb (list_comb (t, ts)) = (t, ts)"
       
    69 *}
       
    70 
       
    71 section {* Some surprises *}
       
    72 
       
    73 ML_command {*
       
    74 check_property "ALL Ts t. type_of1 (Ts, t) = fastype_of1 (Ts, t)"
       
    75 *}
       
    76 
       
    77 
       
    78 ML_command {*
       
    79 val thy = @{theory};
       
    80 check_property "ALL t u. if Pattern.matches thy (t, u) then Term.could_unify (t, u) else true"
       
    81 *}
       
    82 
       
    83 end