src/HOL/List.thy
changeset 28663 bd8438543bf2
parent 28642 658b598d8af4
child 28708 a1a436f09ec6
equal deleted inserted replaced
28662:64ab5bb68d4c 28663:bd8438543bf2
  3518 
  3518 
  3519 ML {*
  3519 ML {*
  3520 local
  3520 local
  3521 
  3521 
  3522 open Basic_Code_Thingol;
  3522 open Basic_Code_Thingol;
  3523 val nil' = Code_Name.const @{theory} @{const_name Nil};
  3523 
  3524 val cons' = Code_Name.const @{theory} @{const_name Cons};
  3524 fun implode_list (nil', cons') t =
  3525 val char' = Code_Name.const @{theory} @{const_name Char}
       
  3526 val nibbles' = map (Code_Name.const @{theory})
       
  3527    [@{const_name Nibble0}, @{const_name Nibble1},
       
  3528     @{const_name Nibble2}, @{const_name Nibble3},
       
  3529     @{const_name Nibble4}, @{const_name Nibble5},
       
  3530     @{const_name Nibble6}, @{const_name Nibble7},
       
  3531     @{const_name Nibble8}, @{const_name Nibble9},
       
  3532     @{const_name NibbleA}, @{const_name NibbleB},
       
  3533     @{const_name NibbleC}, @{const_name NibbleD},
       
  3534     @{const_name NibbleE}, @{const_name NibbleF}];
       
  3535 
       
  3536 fun implode_list t =
       
  3537   let
  3525   let
  3538     fun dest_cons (IConst (c, _) `$ t1 `$ t2) =
  3526     fun dest_cons (IConst (c, _) `$ t1 `$ t2) =
  3539           if c = cons'
  3527           if c = cons'
  3540           then SOME (t1, t2)
  3528           then SOME (t1, t2)
  3541           else NONE
  3529           else NONE
  3544   in case t'
  3532   in case t'
  3545    of IConst (c, _) => if c = nil' then SOME ts else NONE
  3533    of IConst (c, _) => if c = nil' then SOME ts else NONE
  3546     | _ => NONE
  3534     | _ => NONE
  3547   end;
  3535   end;
  3548 
  3536 
  3549 fun decode_char (IConst (c1, _), IConst (c2, _)) =
  3537 fun decode_char nibbles' (IConst (c1, _), IConst (c2, _)) =
  3550       let
  3538       let
  3551         fun idx c = find_index (curry (op =) c) nibbles';
  3539         fun idx c = find_index (curry (op =) c) nibbles';
  3552         fun decode ~1 _ = NONE
  3540         fun decode ~1 _ = NONE
  3553           | decode _ ~1 = NONE
  3541           | decode _ ~1 = NONE
  3554           | decode n m = SOME (chr (n * 16 + m));
  3542           | decode n m = SOME (chr (n * 16 + m));
  3555       in decode (idx c1) (idx c2) end
  3543       in decode (idx c1) (idx c2) end
  3556   | decode_char _ = NONE;
  3544   | decode_char _ _ = NONE;
  3557 
  3545 
  3558 fun implode_string mk_char mk_string ts =
  3546 fun implode_string (char', nibbles') mk_char mk_string ts =
  3559   let
  3547   let
  3560     fun implode_char (IConst (c, _) `$ t1 `$ t2) =
  3548     fun implode_char (IConst (c, _) `$ t1 `$ t2) =
  3561           if c = char' then decode_char (t1, t2) else NONE
  3549           if c = char' then decode_char nibbles' (t1, t2) else NONE
  3562       | implode_char _ = NONE;
  3550       | implode_char _ = NONE;
  3563     val ts' = map implode_char ts;
  3551     val ts' = map implode_char ts;
  3564   in if forall is_some ts'
  3552   in if forall is_some ts'
  3565     then (SOME o Code_Printer.str o mk_string o implode o map_filter I) ts'
  3553     then (SOME o Code_Printer.str o mk_string o implode o map_filter I) ts'
  3566     else NONE
  3554     else NONE
  3567   end;
  3555   end;
  3568 
  3556 
       
  3557 fun list_names naming = pairself (the o Code_Thingol.lookup_const naming)
       
  3558   (@{const_name Nil}, @{const_name Cons});
       
  3559 fun char_name naming = (the o Code_Thingol.lookup_const naming)
       
  3560   @{const_name Char}
       
  3561 fun nibble_names naming = map (the o Code_Thingol.lookup_const naming)
       
  3562   [@{const_name Nibble0}, @{const_name Nibble1},
       
  3563    @{const_name Nibble2}, @{const_name Nibble3},
       
  3564    @{const_name Nibble4}, @{const_name Nibble5},
       
  3565    @{const_name Nibble6}, @{const_name Nibble7},
       
  3566    @{const_name Nibble8}, @{const_name Nibble9},
       
  3567    @{const_name NibbleA}, @{const_name NibbleB},
       
  3568    @{const_name NibbleC}, @{const_name NibbleD},
       
  3569    @{const_name NibbleE}, @{const_name NibbleF}];
       
  3570 
  3569 fun default_list (target_fxy, target_cons) pr fxy t1 t2 =
  3571 fun default_list (target_fxy, target_cons) pr fxy t1 t2 =
  3570   Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy [
  3572   Code_Printer.brackify_infix (target_fxy, Code_Printer.R) fxy [
  3571     pr (Code_Printer.INFX (target_fxy, Code_Printer.X)) t1,
  3573     pr (Code_Printer.INFX (target_fxy, Code_Printer.X)) t1,
  3572     Code_Printer.str target_cons,
  3574     Code_Printer.str target_cons,
  3573     pr (Code_Printer.INFX (target_fxy, Code_Printer.R)) t2
  3575     pr (Code_Printer.INFX (target_fxy, Code_Printer.R)) t2
  3574   ];
  3576   ];
  3575 
  3577 
  3576 fun pretty_list literals =
  3578 fun pretty_list literals =
  3577   let
  3579   let
  3578     val mk_list = Code_Printer.literal_list literals;
  3580     val mk_list = Code_Printer.literal_list literals;
  3579     fun pretty pr thm pat vars fxy [(t1, _), (t2, _)] =
  3581     fun pretty pr naming thm pat vars fxy [(t1, _), (t2, _)] =
  3580       case Option.map (cons t1) (implode_list t2)
  3582       case Option.map (cons t1) (implode_list (list_names naming) t2)
  3581        of SOME ts => mk_list (map (pr vars Code_Printer.NOBR) ts)
  3583        of SOME ts => mk_list (map (pr vars Code_Printer.NOBR) ts)
  3582         | NONE => default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
  3584         | NONE => default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
  3583   in (2, pretty) end;
  3585   in (2, pretty) end;
  3584 
  3586 
  3585 fun pretty_list_string literals =
  3587 fun pretty_list_string literals =
  3586   let
  3588   let
  3587     val mk_list = Code_Printer.literal_list literals;
  3589     val mk_list = Code_Printer.literal_list literals;
  3588     val mk_char = Code_Printer.literal_char literals;
  3590     val mk_char = Code_Printer.literal_char literals;
  3589     val mk_string = Code_Printer.literal_string literals;
  3591     val mk_string = Code_Printer.literal_string literals;
  3590     fun pretty pr thm pat vars fxy [(t1, _), (t2, _)] =
  3592     fun pretty pr naming thm pat vars fxy [(t1, _), (t2, _)] =
  3591       case Option.map (cons t1) (implode_list t2)
  3593       case Option.map (cons t1) (implode_list (list_names naming) t2)
  3592        of SOME ts => (case implode_string mk_char mk_string ts
  3594        of SOME ts => (case implode_string (char_name naming, nibble_names naming) mk_char mk_string ts
  3593            of SOME p => p
  3595            of SOME p => p
  3594             | NONE => mk_list (map (pr vars Code_Printer.NOBR) ts))
  3596             | NONE => mk_list (map (pr vars Code_Printer.NOBR) ts))
  3595         | NONE => default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
  3597         | NONE => default_list (Code_Printer.infix_cons literals) (pr vars) fxy t1 t2;
  3596   in (2, pretty) end;
  3598   in (2, pretty) end;
  3597 
  3599 
  3598 fun pretty_char literals =
  3600 fun pretty_char literals =
  3599   let
  3601   let
  3600     val mk_char = Code_Printer.literal_char literals;
  3602     val mk_char = Code_Printer.literal_char literals;
  3601     fun pretty _ thm _ _ _ [(t1, _), (t2, _)] =
  3603     fun pretty _ naming thm _ _ _ [(t1, _), (t2, _)] =
  3602       case decode_char (t1, t2)
  3604       case decode_char (nibble_names naming) (t1, t2)
  3603        of SOME c => (Code_Printer.str o mk_char) c
  3605        of SOME c => (Code_Printer.str o mk_char) c
  3604         | NONE => Code_Printer.nerror thm "Illegal character expression";
  3606         | NONE => Code_Printer.nerror thm "Illegal character expression";
  3605   in (2, pretty) end;
  3607   in (2, pretty) end;
  3606 
  3608 
  3607 fun pretty_message literals =
  3609 fun pretty_message literals =
  3608   let
  3610   let
  3609     val mk_char = Code_Printer.literal_char literals;
  3611     val mk_char = Code_Printer.literal_char literals;
  3610     val mk_string = Code_Printer.literal_string literals;
  3612     val mk_string = Code_Printer.literal_string literals;
  3611     fun pretty _ thm _ _ _ [(t, _)] =
  3613     fun pretty _ naming thm _ _ _ [(t, _)] =
  3612       case implode_list t
  3614       case implode_list (list_names naming) t
  3613        of SOME ts => (case implode_string mk_char mk_string ts
  3615        of SOME ts => (case implode_string (char_name naming, nibble_names naming) mk_char mk_string ts
  3614            of SOME p => p
  3616            of SOME p => p
  3615             | NONE => Code_Printer.nerror thm "Illegal message expression")
  3617             | NONE => Code_Printer.nerror thm "Illegal message expression")
  3616         | NONE => Code_Printer.nerror thm "Illegal message expression";
  3618         | NONE => Code_Printer.nerror thm "Illegal message expression";
  3617   in (1, pretty) end;
  3619   in (1, pretty) end;
  3618 
  3620