src/HOL/Codegenerator_Test/Generate_Target_GHC.thy
author paulson <lp15@cam.ac.uk>
Sun, 03 Aug 2025 20:34:24 +0100
changeset 82913 7c870287f04f
parent 81999 513f8fa74c82
permissions -rw-r--r--
New lemmas about improper integrals and other things
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
81999
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
     2
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
     3
section \<open>Test of target-language specific implementations for GHC\<close>
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
     4
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
     5
theory Generate_Target_GHC
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
     6
  imports
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
     7
    "HOL-Codegenerator_Test.Generate_Target_String_Literals"
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
     8
    "HOL-Codegenerator_Test.Generate_Target_Bit_Operations"
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
     9
begin
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
    10
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
    11
test_code Generate_Target_String_Literals.check in GHC
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
    12
test_code Generate_Target_Bit_Operations.check in GHC
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
    13
513f8fa74c82 more explicit tests for non-PolyML SML platforms
haftmann
parents:
diff changeset
    14
end