aboutsummaryrefslogtreecommitdiff
path: root/src/Reflection/SmartCastInterp.v
diff options
context:
space:
mode:
authorGravatar Jason Gross <jgross@mit.edu>2017-02-13 14:39:13 -0500
committerGravatar Jason Gross <jgross@mit.edu>2017-02-13 14:39:13 -0500
commitd38d1a272999cf2a4da03d633e319d5d5a4a2779 (patch)
treebb1a5061dbf89857e31ff26239609f23262b68a8 /src/Reflection/SmartCastInterp.v
parent093834a3e35d86ce768102d1c5e894565a36fd74 (diff)
Add SmartCast{Interp,Wf}
Diffstat (limited to 'src/Reflection/SmartCastInterp.v')
-rw-r--r--src/Reflection/SmartCastInterp.v35
1 files changed, 35 insertions, 0 deletions
diff --git a/src/Reflection/SmartCastInterp.v b/src/Reflection/SmartCastInterp.v
new file mode 100644
index 000000000..f7d1a1aef
--- /dev/null
+++ b/src/Reflection/SmartCastInterp.v
@@ -0,0 +1,35 @@
+Require Import Crypto.Reflection.Syntax.
+Require Import Crypto.Reflection.SmartMap.
+Require Import Crypto.Reflection.TypeUtil.
+Require Import Crypto.Reflection.SmartCast.
+Require Import Crypto.Util.Notations.
+
+Local Open Scope expr_scope.
+Local Open Scope ctype_scope.
+Section language.
+ Context {base_type_code : Type}
+ {op : flat_type base_type_code -> flat_type base_type_code -> Type}
+ {interp_base_type : base_type_code -> Type}
+ {interp_op : forall src dst, op src dst -> interp_flat_type interp_base_type src -> interp_flat_type interp_base_type dst}
+ (base_type_beq : base_type_code -> base_type_code -> bool)
+ (base_type_bl_transparent : forall x y, base_type_beq x y = true -> x = y)
+ (Cast : forall var A A', exprf base_type_code op (var:=var) (Tbase A) -> exprf base_type_code op (var:=var) (Tbase A'))
+ (interpf_Cast_id : forall A x, interpf interp_op (Cast _ A A x) = interpf interp_op x)
+ (cast_val : forall A A', interp_base_type A -> interp_base_type A')
+ (interpf_cast : forall A A' e, interpf interp_op (Cast _ A A' e) = cast_val A A' (interpf interp_op e)).
+
+ Local Notation exprf := (@exprf base_type_code op).
+ Local Notation SmartCast_base := (@SmartCast_base _ _ _ base_type_bl_transparent Cast).
+ Local Notation SmartCast := (@SmartCast _ _ _ base_type_bl_transparent Cast).
+
+ Lemma interpf_SmartCast_base {A A'} (x : exprf (Tbase A))
+ : interpf interp_op (SmartCast_base x) = interpf interp_op (Cast _ A A' x).
+ Proof.
+ clear dependent cast_val.
+ unfold SmartCast_base.
+ destruct (Sumbool.sumbool_of_bool (base_type_beq A A')) as [H|H].
+ { destruct (base_type_bl_transparent A A' H).
+ rewrite interpf_Cast_id; reflexivity. }
+ { reflexivity. }
+ Qed.
+End language.