diff --git a/DeepSEA language reference.pdf b/DeepSEA language reference.pdf index eab5f63..b654d1e 100644 Binary files a/DeepSEA language reference.pdf and b/DeepSEA language reference.pdf differ diff --git a/binaries/MacOS/dsc b/binaries/MacOS/dsc index 5e9694a..4b53a22 100755 Binary files a/binaries/MacOS/dsc and b/binaries/MacOS/dsc differ diff --git a/binaries/MacOS/dsc_antchain b/binaries/MacOS/dsc_antchain index 01c03f1..2344c32 100755 Binary files a/binaries/MacOS/dsc_antchain and b/binaries/MacOS/dsc_antchain differ diff --git a/binaries/MacOS/minicc b/binaries/MacOS/minicc new file mode 100755 index 0000000..6d04b2d Binary files /dev/null and b/binaries/MacOS/minicc differ diff --git a/binaries/linux/dsc b/binaries/linux/dsc old mode 100644 new mode 100755 index 62f583b..f7d7d65 Binary files a/binaries/linux/dsc and b/binaries/linux/dsc differ diff --git a/binaries/linux/dsc_antchain b/binaries/linux/dsc_antchain old mode 100644 new mode 100755 index 62f583b..097b00a Binary files a/binaries/linux/dsc_antchain and b/binaries/linux/dsc_antchain differ diff --git a/binaries/linux/minicc b/binaries/linux/minicc new file mode 100755 index 0000000..ecb867f Binary files /dev/null and b/binaries/linux/minicc differ diff --git a/contracts/amm/amm/RefineAMM.v b/contracts/amm/amm/RefineAMM.v new file mode 100644 index 0000000..4137973 --- /dev/null +++ b/contracts/amm/amm/RefineAMM.v @@ -0,0 +1,259 @@ +(* WARNING: This file is generated by Edsger, the DeepSEA compiler. + All modification will be lost when regenerating. *) +(* Module amm.RefineAMM for amm.ds *) +Require Import BinPos. +Require Import DeepSpec.Runtime. +Require Import DeepSpec.Linking. +Require Import amm.EdsgerIdents. +Require Import amm.DataTypes. +Require Import amm.DataTypeOps. +Require Import amm.DataTypeProofs. +Require Import layerlib.LayerCalculusLemma. +Require Import layerlib.RefinementTactic. +Require Import layerlib.RefinementTacticAux. +Require Import liblayers.compcertx.MakeProgram. +Require Import liblayers.compcertx.MemWithData. + +Require Import amm.LayerAMM. +Require Import amm.LSimAMMLIB. + +Section EdsgerGen. + + +Context {mem}`{Hmem: Mem.MemoryModel mem}. +Context`{Hmwd: UseMemWithData mem}. +Context`{make_program_ops: !MakeProgramOps Clight.function Ctypes.type Clight.fundef Ctypes.type}. +Context`{Hmake_program: !MakeProgram Clight.function Ctypes.type Clight.fundef Ctypes.type}. +Instance GlobalLayerSpec : LayerSpecClass := { + make_program_ops := make_program_ops; + Hmake_program := Hmake_program; + GetHighData := global_abstract_data_type +}. +Context`{global_abdata : !GlobalAbData init_global_abstract_data global_low_level_invariant}. +Context`{CTXT_prf : !Layer_AMM_Context_prf}. +Context`{AMMLIB_pres_inv : !AMMLIB_preserves_invariants}. +Context`{AMM_pres_inv : !AMM_preserves_invariants}. + +Existing Instances AMM_overlay_spec AMM_underlay_spec. + +Record relate_RData (j : meminj) (habd : GetHighData) (labd : GetLowData) : Prop + := mkrelate_RData { + (* FixedSupplyToken *) + FixedSupplyToken__totalSupply_re : ; + FixedSupplyToken_balances_re : ; + FixedSupplyToken_allowances_re : ; + (* FixedSupplyToken1 *) + FixedSupplyToken1__totalSupply_re : ; + FixedSupplyToken1_balances_re : ; + FixedSupplyToken1_allowances_re : ; + (* LiquidityToken *) + LiquidityToken__totalSupply_re : ; + LiquidityToken_balances_re : ; + LiquidityToken_allowances_re : +}. + +Record match_RData (habd : GetHighData) (m : mem) (j : meminj) : Prop + := MATCH_RDATA { + AutomatedMarketMaker__token0_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker__token0_var habd m; + AutomatedMarketMaker__token1_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker__token1_var habd m; + AutomatedMarketMaker__owner_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker__owner_var habd m; + AutomatedMarketMaker__reserve0_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker__reserve0_var habd m; + AutomatedMarketMaker__reserve1_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker__reserve1_var habd m; + AutomatedMarketMaker_blockTimestampLast_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker_blockTimestampLast_var habd m; + AutomatedMarketMaker_price0CumulativeLast_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker_price0CumulativeLast_var habd m; + AutomatedMarketMaker_price1CumulativeLast_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker_price1CumulativeLast_var habd m; + AutomatedMarketMaker_kLast_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker_kLast_var habd m +}. + +Local Hint Resolve MATCH_RDATA. + +Global Instance rel_ops: CompatRelOps GetHighDataX GetLowDataX := +{ + relate_AbData f d1 d2 := relate_RData f d1 d2; + match_AbData d1 m f := match_RData d1 m f; + new_glbl := var_AutomatedMarketMaker_AutomatedMarketMaker__token0_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker__token1_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker__owner_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker__reserve0_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker__reserve1_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker_blockTimestampLast_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker_price0CumulativeLast_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker_price1CumulativeLast_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker_kLast_ident :: nil +}. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker__token0_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker__token0_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker__token0_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker__token1_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker__token1_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker__token1_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker__owner_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker__owner_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker__owner_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker__reserve0_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker__reserve0_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker__reserve0_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker__reserve1_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker__reserve1_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker__reserve1_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker_blockTimestampLast_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker_blockTimestampLast_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker_blockTimestampLast_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker_price0CumulativeLast_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker_price0CumulativeLast_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker_price0CumulativeLast_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker_price1CumulativeLast_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker_price1CumulativeLast_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker_price1CumulativeLast_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker_kLast_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker_kLast_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker_kLast_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Lemma relate_incr: + forall abd abd' f f', + relate_RData f abd abd' -> + inject_incr f f' -> + relate_RData f' abd abd'. +Proof. + inversion 1; subst; intros; simpl in *. + repeat match goal with + | H : _ /\ _ |- _ => destruct H + end. + repeat (constructor; simpl; eauto). +Qed. +Global Instance rel_prf: CompatRel GetHighDataX GetLowDataX. +Proof. + constructor; [ destruct 1 .. |]; intros. + - constructor; eapply inject_match_correct; eauto with typeclass_instances. + - constructor; eapply store_match_correct; eauto with typeclass_instances. + - constructor; eapply alloc_match_correct; eauto with typeclass_instances. + - constructor; eapply free_match_correct; eauto with typeclass_instances. + - constructor; eapply storebytes_match_correct; eauto with typeclass_instances. + - eapply relate_incr; eauto. +Qed. + +(* +Local Instance: ExternalCallsOps (mwd GetLowDataX) := CompatExternalCalls.compatlayer_extcall_ops AMMLIB_Layer. +Local Instance: CompilerConfigOps _ := CompatExternalCalls.compatlayer_compiler_config_ops AMMLIB_Layer. +*) + +Instance AMM_hypermem : MemoryModel.HyperMem + := { Hcompatrel := {| crel_prf := rel_prf |} }. + +Class AMM_Underlay_preserves_invariants := { + AMM_Underlay_FixedSupplyToken_constructor_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken_constructor_spec | 5; + AMM_Underlay_FixedSupplyToken_balanceOf_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken_balanceOf_spec | 5; + AMM_Underlay_FixedSupplyToken_transfer_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken_transfer_spec | 5; + AMM_Underlay_FixedSupplyToken_approve_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken_approve_spec | 5; + AMM_Underlay_FixedSupplyToken_transferFrom_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken_transferFrom_spec | 5; + AMM_Underlay_FixedSupplyToken1_constructor_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken1_constructor_spec | 5; + AMM_Underlay_FixedSupplyToken1_balanceOf_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken1_balanceOf_spec | 5; + AMM_Underlay_FixedSupplyToken1_transfer_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken1_transfer_spec | 5; + AMM_Underlay_FixedSupplyToken1_approve_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken1_approve_spec | 5; + AMM_Underlay_FixedSupplyToken1_transferFrom_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken1_transferFrom_spec | 5; + AMM_Underlay_LiquidityToken_constructor_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_constructor_spec | 5; + AMM_Underlay_LiquidityToken_mint_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_mint_spec | 5; + AMM_Underlay_LiquidityToken_burn_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_burn_spec | 5; + AMM_Underlay_LiquidityToken_totalSupply_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_totalSupply_spec | 5; + AMM_Underlay_LiquidityToken_balanceOf_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_balanceOf_spec | 5; + AMM_Underlay_LiquidityToken_transfer_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_transfer_spec | 5; + AMM_Underlay_LiquidityToken_approve_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_approve_spec | 5; + AMM_Underlay_LiquidityToken_transferFrom_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_transferFrom_spec | 5; + AMM_Underlay_AutomatedMarketMakerLib_constructor_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) AutomatedMarketMakerLib_constructor_spec | 5; + AMM_Underlay_AutomatedMarketMakerLib_getAmountIn_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) AutomatedMarketMakerLib_getAmountIn_spec | 5; + AMM_Underlay_AutomatedMarketMakerLib_getBalanceAdjusted_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) AutomatedMarketMakerLib_getBalanceAdjusted_spec | 5; + AMM_Underlay_AutomatedMarketMakerLib_min_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) AutomatedMarketMakerLib_min_spec | 5 +}. +Instance AMM'AMMLIB_preserves_invariants : AMM_Underlay_preserves_invariants. +Proof. esplit; apply AMMLIB_pres_inv. Defined. + +(* +Lemma passthrough_correct: + sim (crel (CompatRel0 := rel_prf) _ _) AMM_Layer_passthrough AMMLIB_Layer. +Proof. + Local Opaque simRR mapsto layer_mapsto_primitive. + unfold GlobalLayerSpec, MemoryModel.GetHighDataX. + simpl. + + sim_oplus; simpl. + + Local Transparent simRR mapsto layer_mapsto_primitive. +Qed.*) +End EdsgerGen. diff --git a/contracts/amm/amm/prf.v b/contracts/amm/amm/prf.v new file mode 100644 index 0000000..aad3bf6 --- /dev/null +++ b/contracts/amm/amm/prf.v @@ -0,0 +1,436 @@ +Require Import amm.DataTypes. +Require Import amm.DataTypeOps. +Require Import lib.Monad.StateMonadOption. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import ZArith. +Require Import core.HyperTypeInst. +Require Import backend.MachineModel. +Require Import amm.LayerAMMLIB. +Require Import amm.LayerAMM. + +Definition state := global_abstract_data_type. + +Definition init_state := init_global_abstract_data. + +Definition wei := Z. + +Definition addr := int256. + +Definition blocknumber := int256. + +Existing Instance GlobalLayerSpec. + +Section step. + + Context (contract_address : addr). + + Section mstep. + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Definition make_machine_env (caller: addr) + : machine_env state + := {| me_address := contract_address; + me_origin := caller; + me_caller := caller; (* need update after every control-flow transfer *) + me_callvalue := Int256.repr (0); + me_coinbase := coinbase; + me_timestamp := timestamp; + me_number := number; + me_balance := balance; + me_blockhash := blockhash; + (* not implemented *) + me_transfer _ _ _ _ _ := False; + me_callmethod _ _ _ _ _ _ _ _ _ _ := False; + me_log _ _ _ := prev_contract_state; + |}. + + Import MonadNotation. + + Definition lif {A:Type} + (caller : addr) + (cmd : machine_env state -> osT global_abstract_data_type A) + : osT state A := + st <- get;; + let me := make_machine_env caller in + match runStateT (cmd me) st with + | None => mzero + | Some (v, st') => put st' ;; ret v + end. + + (* osT state int256 = + state transformer with option monad, state is state, and value is int256 *) + Print osT. (* = fun D : Type => stateT D option *) + Print stateT. (* (S : Type) (m : Type -> Type) (t : Type) : Type := mkStateT + { runStateT : S -> m (t * S)%type } *) + + Print runStateT. (* takes monad transformer, state, returns m (t * S) *) + + (* How the state of the system changes in a single method call made by player p. *) + Inductive mstep u (st st' : state) : Prop := + | FixedSupplyToken_constructor_step : forall r , runStateT (FixedSupplyToken_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_totalSupply_step : forall r , runStateT (FixedSupplyToken_totalSupply_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_balanceOf_step : forall r tokenOwner, runStateT (FixedSupplyToken_balanceOf_opt tokenOwner (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_transfer_step : forall r toA tokens, runStateT (FixedSupplyToken_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_approve_step : forall r spender tokens, runStateT (FixedSupplyToken_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_transferFrom_step : forall r fromA toA tokens, runStateT (FixedSupplyToken_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_constructor_step : forall r , runStateT (FixedSupplyToken1_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_totalSupply_step : forall r , runStateT (FixedSupplyToken1_totalSupply_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_balanceOf_step : forall r tokenOwner, runStateT (FixedSupplyToken1_balanceOf_opt tokenOwner (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_transfer_step : forall r toA tokens, runStateT (FixedSupplyToken1_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_approve_step : forall r spender tokens, runStateT (FixedSupplyToken1_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_transferFrom_step : forall r fromA toA tokens, runStateT (FixedSupplyToken1_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_constructor_step : forall r , runStateT (LiquidityToken_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_mint_step : forall r toA value, runStateT (LiquidityToken_mint_opt toA value (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_burn_step : forall r fromA value, runStateT (LiquidityToken_burn_opt fromA value (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_totalSupply_step : forall r , runStateT (LiquidityToken_totalSupply_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_balanceOf_step : forall r tokenOwner, runStateT (LiquidityToken_balanceOf_opt tokenOwner (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_transfer_step : forall r toA tokens, runStateT (LiquidityToken_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_approve_step : forall r spender tokens, runStateT (LiquidityToken_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_transferFrom_step : forall r fromA toA tokens, runStateT (LiquidityToken_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMakerLib_constructor_step : forall r , runStateT (AutomatedMarketMakerLib_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMakerLib_getAmountIn_step : forall r balance amountOut reserve, runStateT (AutomatedMarketMakerLib_getAmountIn_opt balance amountOut reserve (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMakerLib_getBalanceAdjusted_step : forall r balance amountIn, runStateT (AutomatedMarketMakerLib_getBalanceAdjusted_opt balance amountIn (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMakerLib_min_step : forall r x y, runStateT (AutomatedMarketMakerLib_min_opt x y (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_constructor_step : forall r , runStateT (AutomatedMarketMaker_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_mint_step : forall r toA, runStateT (AutomatedMarketMaker_mint_opt toA (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_burn_step : forall r toA, runStateT (AutomatedMarketMaker_burn_opt toA (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_simpleSwap0_step : forall r toA, runStateT (AutomatedMarketMaker_simpleSwap0_opt toA (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_swap_step : forall r amount0Out amount1Out toA, runStateT (AutomatedMarketMaker_swap_opt amount0Out amount1Out toA (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_skim_step : forall r toA, runStateT (AutomatedMarketMaker_skim_opt toA (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_sync_step : forall r , runStateT (AutomatedMarketMaker_sync_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_k_step : forall r , runStateT (AutomatedMarketMaker_k_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_quote0_step : forall r amount0, runStateT (AutomatedMarketMaker_quote0_opt amount0 (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_getAmountOut0_step : forall r amount0In, runStateT (AutomatedMarketMaker_getAmountOut0_opt amount0In (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_getAmountIn0_step : forall r amount0Out, runStateT (AutomatedMarketMaker_getAmountIn0_opt amount0Out (make_machine_env u)) st = Some (r, st') -> mstep u st st' +. + + (* We can compute a new block by letting the players call the contract + in some arbitrary order. *) + Inductive multi_mstep : state -> state -> Prop := + | multi_mstep_reflexive : forall (st : state), multi_mstep st st + | multi_mstep_transitive : forall (st st' st'' : state) u, + multi_mstep st st' -> mstep u st' st'' -> multi_mstep st st''. + + + (* A block is sufficiently synchronous if every player got a chance to submit a + transaction to to. *) + (* TODO, I think this definition is fine for now, but it seems little too clever, + should probably re-state this using some straightforward tracking of the states + we pass through. *) + Definition multi_mstep_synchronous st1 st2 := + forall u, exists st st', + multi_mstep st1 st /\ mstep u st st' /\ multi_mstep st' st2. + + (* Here are a bunch of induction principles inspired by linear temporal logic. *) + + (* Prove that some property P holds "globally", i.e. for each state along a + path. + You can also specify a property Pprev which is known to hold for the + prev_contract_state. If not needed, it can just be True. + *) + Lemma multi_mstep_Global_ind : forall (Pprev P : state -> Prop), + (forall u st st', Pprev prev_contract_state -> P st -> mstep u st st' -> P st') -> + Pprev prev_contract_state -> forall st st', P st -> multi_mstep st st' -> P st'. + Proof. + induction 4; eauto. + Qed. + + (* + (* Prove that P holds "until" Q along a path. + "Until" is a liveness assertion, so we need the synchronicity assumption. *) + Lemma multi_mstep_Until_ind : forall (Pprev P Q : state -> Prop), + (forall p st st', Pprev prev_contract_state -> P st -> In p players + -> mstep p st st' -> (P st' \/ Q st')) -> + Pprev prev_contract_state -> + forall st, + P st -> exists st', multi_mstep st st' -> (P st' \/ Q st'). + Proof. + induction 4; eauto *) + + End mstep. + + + Definition Int256_incr x := Int256.add x Int256.one. + + Inductive bstep (n : blocknumber) : state -> state -> Prop := + | bstep_step : forall coinbase timestamp balance blockhash st st', + multi_mstep coinbase timestamp n balance blockhash st st st' -> + bstep n st st'. + + Inductive multi_bstep : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_reflexive : forall n (st : state), multi_bstep n st n st + | multi_bstep_transitive : forall n n' (st st' st'' : state), + multi_bstep n st n' st' -> bstep (Int256_incr n') st' st'' -> multi_bstep n st (Int256_incr n') st''. + + (* multi_bstep is the step relation without any synchronicity assumption. + This is sufficient to prove some safety properties, but for most interesting + theorems we instead need to use this synchronous version: *) + + Inductive bstep_synch (n : blocknumber) : state -> state -> Prop := + | bstep_synch_step : forall coinbase timestamp balance blockhash st st', + multi_mstep_synchronous coinbase timestamp n balance blockhash st st st' -> + bstep_synch n st st'. + + Inductive multi_bstep_synch : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_synch_reflexive : forall n (st : state), multi_bstep_synch n st n st + | multi_bstep_synch_transitive : forall n n' (st st' st'' : state), + multi_bstep_synch n st n' st' -> bstep_synch (Int256_incr n') st' st'' -> multi_bstep_synch n st (Int256_incr n') st''. + + Lemma multi_bstep_Global_ind : forall (P : state -> Prop), + (forall coinbase timestamp number balance blockhash prev_block u st st', + P prev_block + -> P st + -> mstep coinbase timestamp number balance blockhash prev_block u st st' + -> P st') + -> forall n st n' st', + P st -> multi_bstep n st n' st' -> P st'. + Proof. + induction 3. + - auto. + - inversion H2; subst. + eapply multi_mstep_Global_ind with (st:=st') (prev_contract_state := st'). + + intros. + refine (H _ _ _ _ _ _ _ _ _ _ _ H6); auto. + + apply IHmulti_bstep; assumption. + + apply IHmulti_bstep; assumption. + + exact H3. + Qed. + End step. + + Section DeepSEAGenericProof. + + Lemma Int256Tree_reduce : forall (i: int256) (v: Z) (t: Int256Tree.t Z), Int256Tree.get_default 0%Z i (Int256Tree.set i v t) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gss . + reflexivity. + Qed. + + Lemma Int256Tree_mreduce : forall (i j : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set i v t)) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gss. + reflexivity. + exact H. + Qed. + + Lemma Int256Tree_mireduce : forall (i j k : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + i <> k -> + j <> k -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set k v t)) = + Int256Tree.get_default 0%Z i t. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gso. + reflexivity. + exact H0. + exact H. + Qed. + + Lemma add_sub_inv : forall (i j : Z32), (i + j - i)%Z = j. + Proof. + intros. + omega. + Qed. + End DeepSEAGenericProof. + + Section Proof. + Context (* (strategies : strategy_tuple) *) + (initial_balances : addr -> Z) + (contract_address : int256). + + Context (init_bt init_rt : int256) + (init_coinbase : int256) + (init_timestamp : int256) + (init_number : int256) + (init_blockhash : int256 -> int256) + (pre_init_state init_state : state). + + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Require Import lib.Monad.RunStateTInv. + Require Import lib.ArithInv. + + Definition make_machine_env_wrapped prev_st user := + make_machine_env contract_address coinbase timestamp number balance blockhash prev_st user. + + Lemma make_machine_env_caller_eq : forall st caller, me_caller (make_machine_env_wrapped st caller) = caller. + Proof. auto. Qed. + + Lemma make_machine_env_address_eq : forall st caller, me_address (make_machine_env_wrapped st caller) = contract_address. + Proof. auto. Qed. + +Transparent FixedSupplyToken_constructor_opt. +Transparent FixedSupplyToken_totalSupply_opt. +Transparent FixedSupplyToken_balanceOf_opt. +Transparent FixedSupplyToken_transfer_opt. +Transparent FixedSupplyToken_approve_opt. +Transparent FixedSupplyToken_transferFrom_opt. +Transparent FixedSupplyToken1_constructor_opt. +Transparent FixedSupplyToken1_totalSupply_opt. +Transparent FixedSupplyToken1_balanceOf_opt. +Transparent FixedSupplyToken1_transfer_opt. +Transparent FixedSupplyToken1_approve_opt. +Transparent FixedSupplyToken1_transferFrom_opt. +Transparent LiquidityToken_constructor_opt. +Transparent LiquidityToken_mint_opt. +Transparent LiquidityToken_burn_opt. +Transparent LiquidityToken_totalSupply_opt. +Transparent LiquidityToken_balanceOf_opt. +Transparent LiquidityToken_transfer_opt. +Transparent LiquidityToken_approve_opt. +Transparent LiquidityToken_transferFrom_opt. +Transparent AutomatedMarketMakerLib_constructor_opt. +Transparent AutomatedMarketMakerLib_getAmountIn_opt. +Transparent AutomatedMarketMakerLib_getBalanceAdjusted_opt. +Transparent AutomatedMarketMakerLib_min_opt. +Transparent AutomatedMarketMaker_constructor_opt. +Transparent AutomatedMarketMaker_mint_opt. +Transparent AutomatedMarketMaker_burn_opt. +Transparent AutomatedMarketMaker_simpleSwap0_opt. +Transparent AutomatedMarketMaker_swap_opt. +Transparent AutomatedMarketMaker_skim_opt. +Transparent AutomatedMarketMaker_sync_opt. +Transparent AutomatedMarketMaker_k_opt. +Transparent AutomatedMarketMaker_quote0_opt. +Transparent AutomatedMarketMaker_getAmountOut0_opt. +Transparent AutomatedMarketMaker_getAmountIn0_opt. +Definition FixedSupplyToken_address := (Int256.repr 65587). +Definition FixedSupplyToken1_address := (Int256.repr 65586). +Definition LiquidityToken_address := (Int256.repr 65585). +Definition AutomatedMarketMaker_address := (Int256.repr 65584). + +Definition not_contract_address (addr: int256) := + addr <> FixedSupplyToken_address /\ + addr <> FixedSupplyToken1_address /\ + addr <> LiquidityToken_address /\ + addr <> AutomatedMarketMaker_address. + +Axiom taddr : contract_address = AutomatedMarketMaker_address. + +Ltac rds := +unfold FixedSupplyToken_constructor_opt in *; +unfold FixedSupplyToken_totalSupply_opt in *; +unfold FixedSupplyToken_balanceOf_opt in *; +unfold FixedSupplyToken_transfer_opt in *; +unfold FixedSupplyToken_approve_opt in *; +unfold FixedSupplyToken_transferFrom_opt in *; +unfold FixedSupplyToken1_constructor_opt in *; +unfold FixedSupplyToken1_totalSupply_opt in *; +unfold FixedSupplyToken1_balanceOf_opt in *; +unfold FixedSupplyToken1_transfer_opt in *; +unfold FixedSupplyToken1_approve_opt in *; +unfold FixedSupplyToken1_transferFrom_opt in *; +unfold LiquidityToken_constructor_opt in *; +unfold LiquidityToken_mint_opt in *; +unfold LiquidityToken_burn_opt in *; +unfold LiquidityToken_totalSupply_opt in *; +unfold LiquidityToken_balanceOf_opt in *; +unfold LiquidityToken_transfer_opt in *; +unfold LiquidityToken_approve_opt in *; +unfold LiquidityToken_transferFrom_opt in *; +unfold AutomatedMarketMakerLib_constructor_opt in *; +unfold AutomatedMarketMakerLib_getAmountIn_opt in *; +unfold AutomatedMarketMakerLib_getBalanceAdjusted_opt in *; +unfold AutomatedMarketMakerLib_min_opt in *; +unfold AutomatedMarketMaker_constructor_opt in *; +unfold AutomatedMarketMaker_mint_opt in *; +unfold AutomatedMarketMaker_burn_opt in *; +unfold AutomatedMarketMaker_simpleSwap0_opt in *; +unfold AutomatedMarketMaker_swap_opt in *; +unfold AutomatedMarketMaker_skim_opt in *; +unfold AutomatedMarketMaker_sync_opt in *; +unfold AutomatedMarketMaker_k_opt in *; +unfold AutomatedMarketMaker_quote0_opt in *; +unfold AutomatedMarketMaker_getAmountOut0_opt in *; +unfold AutomatedMarketMaker_getAmountIn0_opt in *; +inv_runStateT; +subst; +inv_arith; +simpl; +try rewrite make_machine_env_caller_eq in *; +try rewrite make_machine_env_address_eq in *; +try rewrite Int256Tree_reduce in *; +try rewrite Int256Tree_mreduce in *; +try rewrite Int256Tree_mireduce in *; +auto. + + Definition compute_k (s: state) := + Z.mul (AutomatedMarketMaker__reserve0 s) (AutomatedMarketMaker__reserve1 s). + + Definition get_balance0 (s: state) (a: addr) := + (Int256Tree.get_default (0%Z) + a (FixedSupplyToken_balances s)). + + Definition get_balance1 (s: state) (a: addr) := + (Int256Tree.get_default 0%Z + a (FixedSupplyToken1_balances s)). + + (* need to transfer balance first, assume token0 transfer first *) + (* should have been approve and then transferFrom, here we are hacking a bit *) + (* prove that after every possible function call, whichever layers, the + k = _reserve0 * _reserve1 is strictly increasing for simpleSwap0: + + newState = simpleSwap0 oldState /\ compute_k newState > compute_k oldState *) + Theorem increasing_k_simpleSwap0 : (forall r r' + (s s' s'' : state) + (trader : addr) + (swapAmount0 : Z32), + (swapAmount0 > 0)%Z -> (* trasferring an amount > 0 *) + ((AutomatedMarketMaker__reserve0 s) > 0)%Z -> + ((AutomatedMarketMaker__reserve1 s) > 0)%Z -> + not_contract_address trader -> + get_balance0 s AutomatedMarketMaker_address = (AutomatedMarketMaker__reserve0 s) -> (* assumes that the reserve is the same as balance at s *) + get_balance1 s AutomatedMarketMaker_address = (AutomatedMarketMaker__reserve1 s) -> + runStateT (FixedSupplyToken_transfer_opt AutomatedMarketMaker_address swapAmount0 (make_machine_env_wrapped s trader)) s = Some (r, s') -> + runStateT (AutomatedMarketMaker_simpleSwap0_opt trader (make_machine_env_wrapped s trader)) s' = Some (r', s'') -> + Z.lt (compute_k s) (compute_k s'')). + Proof. + intros. + unfold get_balance0 in *. + unfold get_balance1 in *. + unfold compute_k. + unfold AutomatedMarketMaker_simpleSwap0_opt in *. + unfold FixedSupplyToken1_transfer_opt in *. + unfold FixedSupplyToken_balanceOf_opt in *. + unfold FixedSupplyToken1_balanceOf_opt in *. + unfold FixedSupplyToken_transfer_opt in *. + inv_runStateT. + rds. + + repeat match goal with + | [ |- context[AutomatedMarketMaker__reserve1 ?X]] => remember (AutomatedMarketMaker__reserve1 X) as R1 + end. + repeat match goal with + | [ |- context[AutomatedMarketMaker__reserve0 ?X]] => remember (AutomatedMarketMaker__reserve0 X) as R0 + end. + match goal with + | [ |- context[Z.mul (Z.add ?X _) _]] => replace X with R0 (* rewrite -> H1 *) + end. + match goal with + | [ |- context[Z.sub ?X _]] => replace X with R1 + end. + remember swapAmount0 as delta0. + +End Proof. diff --git a/contracts/amm/amm/prf_int.v b/contracts/amm/amm/prf_int.v new file mode 100644 index 0000000..484022b --- /dev/null +++ b/contracts/amm/amm/prf_int.v @@ -0,0 +1,631 @@ +Require Import amm.DataTypes. +Require Import amm.DataTypeOps. +Require Import amm.LayerAMM. +Require Import lib.Monad.StateMonadOption. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import ZArith. +Require Import core.HyperTypeInst. +(* for `hashvalue`, maybe should move that to some different file. *) +Require Import backend.MachineModel. + +(* PROOFS: + 2. balance tracks reserve + 3. increasing k -> algebra + 4. no path independence, splitting trades is strictly more expensive -> algebra, appendix D + 5. liquidity token economics is right, i.e., UNI-V2 tracks the actual share of pool --> actually, where is the trading of UNI-V2 Pair tokens? What is its value? + 6. MOD: slippage control + 7. ... cost of manipulation (hard) +*) + +(* global_abstract_data_type *) + +Definition state := global_abstract_data_type. + +Definition init_state := init_global_abstract_data. + +Definition wei := Z. + +Definition addr := int256. + +Definition blocknumber := int256. + +Existing Instance GlobalLayerSpec. + +Section step. + + Context (token0_address token1_address amm_address : addr). + + Section mstep. + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Definition make_machine_env (caller: addr) + : machine_env state + := {| me_address := amm_address; + me_origin := caller; + me_caller := caller; + me_callvalue := Int256.repr (0); + me_coinbase := coinbase; + me_timestamp := timestamp; + me_number := number; + me_balance := balance; + me_blockhash := blockhash; + (* not implemented *) + me_transfer _ _ _ _ _ := False; + me_callmethod _ _ _ _ _ _ _ _ _ _ := False; + me_log _ _ _ := prev_contract_state; + |}. + + Import MonadNotation. + + Definition lif {A:Type} + (caller : addr) + (cmd : machine_env state -> osT global_abstract_data_type A) + : osT state A := + st <- get;; + let me := make_machine_env caller in + match runStateT (cmd me) st with + | None => mzero + | Some (v, st') => put st' ;; ret v + end. + + Definition lifted_simpleSwap0 := + fun (caller: addr) (callvalue: wei) (toA: addr) => lif caller + (AutomatedMarketMaker_simpleSwap0_opt toA). + + Definition lifted_LiquidityToken_mint := + fun (caller: addr) (callvalue: wei) (toA: addr) (value: Z) => lif caller + (LiquidityToken_mint_opt toA value). + + (* osT state int256 = + state transformer with option monad, state is state, and value is int256 *) + Print osT. (* = fun D : Type => stateT D option *) + Print stateT. (* (S : Type) (m : Type -> Type) (t : Type) : Type := mkStateT + { runStateT : S -> m (t * S)%type } *) + + Print runStateT. (* takes monad transformer, state, returns m (t * S) *) + + (* How the state of the system changes in a single method call made by player p. *) + (* for simplicity we only model functions that modify states and can be called by user *) + Inductive mstep u (st st' : state) : Prop := + (* fixed supply tokens *) + | FixedSupplyToken0_transfer_step : + forall toA tokens r, runStateT (FixedSupplyToken_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | FixedSupplyToken1_transfer_step : + forall toA tokens r, runStateT (FixedSupplyToken1_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | FixedSupplyToken0_approve_step : + forall spender tokens r, runStateT (FixedSupplyToken_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | FixedSupplyToken1_approve_step : + forall spender tokens r, runStateT (FixedSupplyToken1_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | FixedSupplyToken0_transferFrom_step : + forall fromA toA tokens r, runStateT (FixedSupplyToken_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | FixedSupplyToken1_transferFrom_step : + forall fromA toA tokens r, runStateT (FixedSupplyToken1_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + (* amm *) + | AutomatedMarketMaker_mint_step : + forall toA r, runStateT (AutomatedMarketMaker_mint_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | AutomatedMarketMaker_burn_step : + forall toA r, runStateT (AutomatedMarketMaker_burn_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | AutomatedMarketMaker_simpleSwap0_step : + forall toA r, runStateT (AutomatedMarketMaker_simpleSwap0_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + . + + (* How the state of the system changes in a single method call made by player p. *) + (* for simplicity we only model functions that modify states and can be called by user *) + Inductive mstep_amm u (st st' : state) : Prop := + (* amm *) + | AutomatedMarketMaker_mint_ammstep : + forall toA r, runStateT (AutomatedMarketMaker_mint_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep_amm u st st' + | AutomatedMarketMaker_burn_ammstep : + forall toA r, runStateT (AutomatedMarketMaker_burn_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep_amm u st st' + | AutomatedMarketMaker_simpleSwap0_ammstep : + forall toA r, runStateT (AutomatedMarketMaker_simpleSwap0_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep_amm u st st' + | AutomatedMarketMaker_skim_ammstep : + forall toA r, runStateT (AutomatedMarketMaker_skim_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep_amm u st st' + | AutomatedMarketMaker_sync_ammstep : + forall r, runStateT (AutomatedMarketMaker_sync_opt (make_machine_env u)) st = Some (r, st') + -> mstep_amm u st st' + . + + (* We can compute a new block by letting the players call the contract + in some arbitrary order. *) + Inductive multi_mstep : state -> state -> Prop := + | multi_mstep_reflexive : forall (st : state), multi_mstep st st + | multi_mstep_transitive : forall (st st' st'' : state) u, + multi_mstep st st' -> mstep u st' st'' -> multi_mstep st st''. + + + (* A block is sufficiently synchronous if every player got a chance to submit a + transaction to to. *) + (* TODO, I think this definition is fine for now, but it seems little too clever, + should probably re-state this using some straightforward tracking of the states + we pass through. *) + Definition multi_mstep_synchronous st1 st2 := + forall u, exists st st', + multi_mstep st1 st /\ mstep u st st' /\ multi_mstep st' st2. + + (* Here are a bunch of induction principles inspired by linear temporal logic. *) + + (* Prove that some property P holds "globally", i.e. for each state along a + path. + You can also specify a property Pprev which is known to hold for the + prev_contract_state. If not needed, it can just be True. + *) + Lemma multi_mstep_Global_ind : forall (Pprev P : state -> Prop), + (forall u st st', Pprev prev_contract_state -> P st -> mstep u st st' -> P st') -> + Pprev prev_contract_state -> forall st st', P st -> multi_mstep st st' -> P st'. + Proof. + induction 4; eauto. + Qed. + + (* + (* Prove that P holds "until" Q along a path. + "Until" is a liveness assertion, so we need the synchronicity assumption. *) + Lemma multi_mstep_Until_ind : forall (Pprev P Q : state -> Prop), + (forall p st st', Pprev prev_contract_state -> P st -> In p players + -> mstep p st st' -> (P st' \/ Q st')) -> + Pprev prev_contract_state -> + forall st, + P st -> exists st', multi_mstep st st' -> (P st' \/ Q st'). + Proof. + induction 4; eauto *) + + End mstep. + + + Definition Int256_incr x := Int256.add x Int256.one. + + Inductive bstep (n : blocknumber) : state -> state -> Prop := + | bstep_step : forall coinbase timestamp balance blockhash st st', + multi_mstep coinbase timestamp n balance blockhash st st st' -> + bstep n st st'. + + Inductive multi_bstep : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_reflexive : forall n (st : state), multi_bstep n st n st + | multi_bstep_transitive : forall n n' (st st' st'' : state), + multi_bstep n st n' st' -> bstep (Int256_incr n') st' st'' -> multi_bstep n st (Int256_incr n') st''. + + (* multi_bstep is the step relation without any synchronicity assumption. + This is sufficient to prove some safety properties, but for most interesting + theorems we instead need to use this synchronous version: *) + + Inductive bstep_synch (n : blocknumber) : state -> state -> Prop := + | bstep_synch_step : forall coinbase timestamp balance blockhash st st', + multi_mstep_synchronous coinbase timestamp n balance blockhash st st st' -> + bstep_synch n st st'. + + Inductive multi_bstep_synch : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_synch_reflexive : forall n (st : state), multi_bstep_synch n st n st + | multi_bstep_synch_transitive : forall n n' (st st' st'' : state), + multi_bstep_synch n st n' st' -> bstep_synch (Int256_incr n') st' st'' -> multi_bstep_synch n st (Int256_incr n') st''. + + Lemma multi_bstep_Global_ind : forall (P : state -> Prop), + (forall coinbase timestamp number balance blockhash prev_block u st st', + P prev_block + -> P st + -> mstep coinbase timestamp number balance blockhash prev_block u st st' + -> P st') + -> forall n st n' st', + P st -> multi_bstep n st n' st' -> P st'. + Proof. + induction 3. + - auto. + - inversion H2; subst. + eapply multi_mstep_Global_ind with (st:=st') (prev_contract_state := st'). + + intros. + refine (H _ _ _ _ _ _ _ _ _ _ _ H6); auto. + + apply IHmulti_bstep; assumption. + + apply IHmulti_bstep; assumption. + + exact H3. + Qed. +End step. + +Section DeepSEAGenericProof. + + Lemma Int256Tree_reduce : forall (i: int256) (v: Z) (t: Int256Tree.t Z), Int256Tree.get_default 0%Z i (Int256Tree.set i v t) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gss . + reflexivity. + Qed. + + Lemma Int256Tree_mreduce : forall (i j : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set i v t)) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gss. + reflexivity. + exact H. + Qed. + + Lemma Int256Tree_mireduce : forall (i j k : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + i <> k -> + j <> k -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set k v t)) = + Int256Tree.get_default 0%Z i t. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gso. + reflexivity. + exact H0. + exact H. + Qed. + + Lemma add_sub_inv : forall (i j : Z32), (i + j - i)%Z = j. + Proof. + intros. + omega. + Qed. + + (* Lemmas on Inequalities over Z *) + Lemma ZInEq_ladd : (forall (i j k : Z32), (i > j + k <-> i - k > j)%Z). + Proof. + intros. omega. + Qed. + Hint Rewrite ZInEq_ladd : ZInEq. + + Axiom ZInEq_multpos : (forall (a b c : Z32), c > 0 -> a > b <-> a * c > b * c)%Z. + Hint Rewrite ZInEq_multpos : ZInEq. + + Axiom ZInEq_ldiv : (forall (a b c : Z32), b > 0 -> a / b > c <-> a > b * c)%Z. + + Axiom ZInEq_denadd : (forall (a b c : Z32), c > 0 -> a / b > a / (b + c))%Z. + + Lemma increasing_k_math : forall (i j k : Z), (i > 0 -> j > 0 -> k > 0 -> + i * j < (i + k) * (j - (k * 997 * j) / (i * 1000 + k * 997)))%Z. + Proof. + From Coq Require Import ZArith Psatz. + Open Scope Z_scope. + simpl. Fail lia. Print Lia.lia. Print Lia.zchecker_no_abstract. + intros. + (* + psatz Z 100. + lia. + + assert (forall (v v' v'' : Z), v'' > 0 -> v < v' -> v * v'' < v' * v''). + intros. + pose (H19 := H16 (v * v'') (v' * v'')); clearbody H19. + apply H19. + unfold Z.sub. + apply mul_sub_distr. + pose (H17 := H16 (i * j) ((i + k) * (j - k * 997 * j / (i * 1000 + k * 997))) (i * 1000 + k * 997)); clearbody H17. + *) + assert (0 < (i * 1000 + k * 997)). omega. + (* multiplied both sides by the denominator of the divisor and needed to use comparisons between "(W / V) * V" and W *) + pose (Zmult_nc:=(Zmult_lt_compat_r (i * j) ((i + k) * (j - k * 997 * j / (i * 1000 + k * 997))) (i * 1000 + k * 997) H2)); clearbody Zmult_nc. + Admitted. +End DeepSEAGenericProof. + +Section AMMProof. + Context (token0_address token1_address amm_address : addr). + + Context (* (strategies : strategy_tuple) *) + (initial_balances : addr -> Z) + (contract_address : int256). + + Context (init_bt init_rt : int256) + (init_coinbase : int256) + (init_timestamp : int256) + (init_number : int256) + (init_blockhash : int256 -> int256) + (pre_init_state init_state : state). + + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Require Import lib.Monad.RunStateTInv. + Require Import lib.ArithInv. + + Definition compute_k (s: state) := + Z.mul (AutomatedMarketMaker__reserve0 s) (AutomatedMarketMaker__reserve1 s). + + Definition get_balance0 (s: state) (a: addr) := + (Int256Tree.get_default (0%Z) + a (FixedSupplyToken_balances s)). + + Definition get_balance1 (s: state) (a: addr) := + (Int256Tree.get_default 0%Z + a (FixedSupplyToken1_balances s)). + + Definition make_machine_env_wrapped prev_st user := + make_machine_env amm_address coinbase timestamp number balance blockhash prev_st user. + + Lemma make_machine_env_caller_eq : forall st caller, me_caller (make_machine_env_wrapped st caller) = caller. + Proof. auto. Qed. + + Lemma make_machine_env_address_eq : forall st caller, me_address (make_machine_env_wrapped st caller) = amm_address. + Proof. auto. Qed. + +Transparent FixedSupplyToken_constructor_opt. +Transparent FixedSupplyToken_totalSupply_opt. +Transparent FixedSupplyToken_balanceOf_opt. +Transparent FixedSupplyToken_transfer_opt. +Transparent FixedSupplyToken_approve_opt. +Transparent FixedSupplyToken_transferFrom_opt. +Transparent FixedSupplyToken1_constructor_opt. +Transparent FixedSupplyToken1_totalSupply_opt. +Transparent FixedSupplyToken1_balanceOf_opt. +Transparent FixedSupplyToken1_transfer_opt. +Transparent FixedSupplyToken1_approve_opt. +Transparent FixedSupplyToken1_transferFrom_opt. +Transparent LiquidityToken_constructor_opt. +Transparent LiquidityToken_mint_opt. +Transparent LiquidityToken_burn_opt. +Transparent LiquidityToken_totalSupply_opt. +Transparent LiquidityToken_balanceOf_opt. +Transparent LiquidityToken_transfer_opt. +Transparent LiquidityToken_approve_opt. +Transparent LiquidityToken_transferFrom_opt. +Transparent AutomatedMarketMakerLib_constructor_opt. +Transparent AutomatedMarketMakerLib_getAmountIn_opt. +Transparent AutomatedMarketMakerLib_getBalanceAdjusted_opt. +Transparent AutomatedMarketMakerLib_min_opt. +Transparent AutomatedMarketMaker_constructor_opt. +Transparent AutomatedMarketMaker_mint_opt. +Transparent AutomatedMarketMaker_burn_opt. +Transparent AutomatedMarketMaker_simpleSwap0_opt. +Transparent AutomatedMarketMaker_swap_opt. +Transparent AutomatedMarketMaker_skim_opt. +Transparent AutomatedMarketMaker_sync_opt. +Transparent AutomatedMarketMaker_k_opt. +Transparent AutomatedMarketMaker_quote0_opt. +Transparent AutomatedMarketMaker_getAmountOut0_opt. +Transparent AutomatedMarketMaker_getAmountIn0_opt. +Definition FixedSupplyToken_address := (Int256.repr 65587). +Definition FixedSupplyToken1_address := (Int256.repr 65586). +Definition LiquidityToken_address := (Int256.repr 65585). +Definition AutomatedMarketMaker_address := (Int256.repr 65584). + +Ltac rds := +unfold FixedSupplyToken_constructor_opt in *; +unfold FixedSupplyToken_totalSupply_opt in *; +unfold FixedSupplyToken_balanceOf_opt in *; +unfold FixedSupplyToken_transfer_opt in *; +unfold FixedSupplyToken_approve_opt in *; +unfold FixedSupplyToken_transferFrom_opt in *; +unfold FixedSupplyToken1_constructor_opt in *; +unfold FixedSupplyToken1_totalSupply_opt in *; +unfold FixedSupplyToken1_balanceOf_opt in *; +unfold FixedSupplyToken1_transfer_opt in *; +unfold FixedSupplyToken1_approve_opt in *; +unfold FixedSupplyToken1_transferFrom_opt in *; +unfold LiquidityToken_constructor_opt in *; +unfold LiquidityToken_mint_opt in *; +unfold LiquidityToken_burn_opt in *; +unfold LiquidityToken_totalSupply_opt in *; +unfold LiquidityToken_balanceOf_opt in *; +unfold LiquidityToken_transfer_opt in *; +unfold LiquidityToken_approve_opt in *; +unfold LiquidityToken_transferFrom_opt in *; +unfold AutomatedMarketMakerLib_constructor_opt in *; +unfold AutomatedMarketMakerLib_getAmountIn_opt in *; +unfold AutomatedMarketMakerLib_getBalanceAdjusted_opt in *; +unfold AutomatedMarketMakerLib_min_opt in *; +unfold AutomatedMarketMaker_constructor_opt in *; +unfold AutomatedMarketMaker_mint_opt in *; +unfold AutomatedMarketMaker_burn_opt in *; +unfold AutomatedMarketMaker_simpleSwap0_opt in *; +unfold AutomatedMarketMaker_swap_opt in *; +unfold AutomatedMarketMaker_skim_opt in *; +unfold AutomatedMarketMaker_sync_opt in *; +unfold AutomatedMarketMaker_k_opt in *; +unfold AutomatedMarketMaker_quote0_opt in *; +unfold AutomatedMarketMaker_getAmountOut0_opt in *; +unfold AutomatedMarketMaker_getAmountIn0_opt in *; +inv_runStateT; +subst; +inv_arith; +simpl in *; +try rewrite make_machine_env_caller_eq in *; +try rewrite make_machine_env_address_eq in *; +try rewrite Int256Tree_reduce in *; +try rewrite Int256Tree_mreduce in *; +try rewrite Int256Tree_mireduce in *; +auto. + + (* need to transfer balance first, assume token0 transfer first *) + (* should have been approve and then transferFrom, here we are hacking a bit *) + (* prove that after every possible function call, whichever layers, the + k = _reserve0 * _reserve1 is strictly increasing for simpleSwap0: + + newState = simpleSwap0 oldState /\ compute_k newState > compute_k oldState *) + Theorem increasing_k_simpleSwap0 : (forall r r' + (s s' s'' : state) + (trader : addr) + (swapAmount0 : Z32), + (swapAmount0 > 0)%Z -> (* trasferring an amount > 0 *) + ((AutomatedMarketMaker__reserve0 s) > 0)%Z -> + ((AutomatedMarketMaker__reserve1 s) > 0)%Z -> + amm_address <> trader -> + get_balance0 s amm_address = (AutomatedMarketMaker__reserve0 s) -> (* assumes that the reserve is the same as balance at s *) + get_balance1 s amm_address = (AutomatedMarketMaker__reserve1 s) -> + runStateT (FixedSupplyToken_transfer_opt amm_address swapAmount0 (make_machine_env_wrapped s trader)) s = Some (r, s') -> + runStateT (AutomatedMarketMaker_simpleSwap0_opt trader (make_machine_env_wrapped s amm_address)) s' = Some (r', s'') -> + Z.lt (compute_k s) (compute_k s'')). + Proof. + intros. + unfold get_balance0 in H0. + unfold get_balance1 in H1. + unfold AutomatedMarketMaker_simpleSwap0_opt in *. + unfold FixedSupplyToken1_transfer_opt in *. + unfold FixedSupplyToken_balanceOf_opt in *. + unfold FixedSupplyToken1_balanceOf_opt in *. + unfold FixedSupplyToken_transfer_opt in *. + inv_runStateT. + rewrite -> make_machine_env_caller_eq in *. + rewrite -> make_machine_env_address_eq in *. + subst. + inv_arith. + unfold compute_k. + simpl. + + rewrite Int256Tree_reduce in *. + rewrite Int256Tree_mreduce in *. + repeat match goal with + | [ |- context[AutomatedMarketMaker__reserve1 ?X]] => remember (AutomatedMarketMaker__reserve1 X) as R1 + end. + repeat match goal with + | [ |- context[AutomatedMarketMaker__reserve0 ?X]] => remember (AutomatedMarketMaker__reserve0 X) as R0 + end. + match goal with + | [ |- context[Z.mul (Z.add ?X _) _]] => replace X with R0 (* rewrite -> H1 *) + end. + match goal with + | [ |- context[Z.sub ?X _]] => replace X with R1 + end. + remember swapAmount0 as delta0. + + pose (IKM := increasing_k_math R0 R1 delta0 H0 H1 H); clearbody IKM. + rewrite add_sub_inv. + exact IKM. + exact H2. + Qed. + +Lemma balance_tracks_reserve_inv : forall u s s', + u <> amm_address -> + get_balance0 s amm_address = (AutomatedMarketMaker__reserve0 s) -> + get_balance1 s amm_address = (AutomatedMarketMaker__reserve1 s) -> + mstep_amm amm_address coinbase timestamp number balance blockhash prev_contract_state u s s' -> + get_balance0 s amm_address = (AutomatedMarketMaker__reserve0 s') /\ + get_balance1 s amm_address = (AutomatedMarketMaker__reserve1 s'). +Proof. + intros u s s' Hun Hps0 Hps1 Hstp. + unfold get_balance0 in *. + unfold get_balance1 in *. + destruct Hstp; rds. (* FIXME: rds should be able to automatically solve all the goals if we have + machine_env updating implemented *) +Admitted. + +Axiom anything : False. + +(* splitting trades is strictly more expensive, if you split one trade into two trades with combined + same amount, then you swapped for less alternative tokens *) +Theorem path_dependency : forall r r' rd rd' rd'' rd''' + (s s' s'' sd' sd'' sd''' sd'''' : state) + (trader : addr) + (swapAmount0 : Z32) + (swapAmount0D0 swapAmount0D1 : Z32), + (swapAmount0D0 > 0)%Z -> + (swapAmount0D1 > 0)%Z -> + (swapAmount0D0 + swapAmount0D1 = swapAmount0)%Z -> + (swapAmount0 > 0)%Z -> (* trasferring an amount > 0 *) + ((AutomatedMarketMaker__reserve0 s) > 0)%Z -> + ((AutomatedMarketMaker__reserve1 s) > 0)%Z -> + amm_address <> trader -> + get_balance0 s amm_address = (AutomatedMarketMaker__reserve0 s) -> (* assumes that the reserve is the same as balance at s *) + get_balance1 s amm_address = (AutomatedMarketMaker__reserve1 s) -> + (* swapping in one trade *) + runStateT (FixedSupplyToken0_transfer_opt amm_address swapAmount0 (make_machine_env_wrapped s trader)) s = Some (r, s') -> + runStateT (AutomatedMarketMaker_simpleSwap0_opt trader (make_machine_env_wrapped s amm_address)) s' = Some (r', s'') -> + (* swapping in two trades *) + runStateT (FixedSupplyToken0_transfer_opt amm_address swapAmount0D0 (make_machine_env_wrapped s trader)) s = Some (rd, sd') -> + runStateT (AutomatedMarketMaker_simpleSwap0_opt trader (make_machine_env_wrapped s amm_address)) sd' = Some (rd', sd'') -> + runStateT (FixedSupplyToken0_transfer_opt amm_address swapAmount0D1 (make_machine_env_wrapped s trader)) sd'' = Some (rd'', sd''') -> + runStateT (AutomatedMarketMaker_simpleSwap0_opt trader (make_machine_env_wrapped s amm_address)) sd''' = Some (rd''', sd'''') -> + (* the ending balance is different *) + (r' > rd' + rd''')%Z. +Proof. + (* generic tactic *) + intros. + unfold get_balance0 in *. + unfold get_balance1 in *. + Time rds. (* finish in 182 secs, or 3m2s *) + rewrite !H6. rewrite !H7. + remember (AutomatedMarketMaker__reserve0 m121) as a. + remember (AutomatedMarketMaker__reserve1 m121) as b. + remember (swapAmount0D0) as c. + remember (swapAmount0D1) as d. + rewrite !add_sub_inv. (* need inequality rewriting rules *) + + assert (d * 997 * R1 / (R0 * 1000 + d * 997) > d * 997 * R1 / (R0 * 1000 + d * 997 + d' * 997))%Z. + assert (d' * 997 > 0)%Z. omega. + match goal with + | [ |- context[Z.gt (Z.div ?A ?B) (Z.div ?A (Z.add ?B ?C))]] => pose (HZInEq_denadd := ZInEq_denadd A B C H31); clearbody HZInEq_denadd + end. + exact HZInEq_denadd. + apply H31. + rewrite Zgt_trans. + nia. + unfold AutomatedMarketMaker_simpleSwap0_opt in *. + unfold FixedSupplyToken1_transfer_opt in *. + unfold FixedSupplyToken0_balanceOf_opt in *. + unfold FixedSupplyToken1_balanceOf_opt in *. + unfold FixedSupplyToken0_transfer_opt in *. + + Opaque bind ret get gets put modify guard mzero. + + Import MonadNotation. + inv_runStateT1 H13. + subst. + inv_arith. + simpl in *. + + inv_runStateT1 H12. + subst. + autorewrite with updates in *. + rewrite !Int256Tree_reduce in *. + simpl in *. + inv_arith. + + inv_runStateT1 H11. + subst. + simpl in *. + autorewrite with updates in *. + inv_arith. + rewrite !(@Int256Tree_Properties.get_default_so _ (0%Z) a25 a) in * by (exact H22). + rewrite !Int256Tree_Properties.get_default_ss in *. + + inv_runStateT1 H9. + subst. + simpl in *. + autorewrite with updates in *. + inv_arith. + + inv_runStateT1 H8. + subst. + simpl in *. + autorewrite with updates in *. + rewrite !Int256Tree_Properties.get_default_ss in *. + + inv_runStateT1 H10. + subst. + simpl in *. + autorewrite with updates in *. + rewrite !Int256Tree_Properties.get_default_ss in *. + + (* These lines are to test if the resulting proof term can be checked reasonably quickly, but it seems it works. *) + (* destruct anything. + Time Qed. *) + +Admitted. diff --git a/contracts/auction/auction.ds b/contracts/auction/auction.ds old mode 100644 new mode 100755 index 11ab494..df99c48 --- a/contracts/auction/auction.ds +++ b/contracts/auction/auction.ds @@ -2,24 +2,22 @@ - implement withdrawals *) -type addr := uint - object signature AuctionInterface = { initialize : uint -> unit; - getbid : unit -> uint; - getbidder : unit -> addr; - getchair : unit -> addr; - getdeadline : unit -> uint; + const getbid : unit -> uint; + const getbidder : unit -> address; + const getchair : unit -> address; + const getdeadline : unit -> uint; bid : unit -> bool; } object OpenAuction : AuctionInterface { (* initial values don't get used *) let _bid : uint := 0u0 - let _bidder : addr := 0u0 - let _chair : addr := 0u0 + let _bidder : address := address(0) + let _chair : address := address(0) let _deadline : uint := 0u0 - let withdrawals : mapping[addr] uint := mapping_init + let withdrawals : mapping[address] uint := mapping_init (* placeholder for constructor *) let initialize deadline = diff --git a/contracts/auction/auction/FunctionalCorrectness.v b/contracts/auction/auction/FunctionalCorrectness.v new file mode 100755 index 0000000..cf0f156 --- /dev/null +++ b/contracts/auction/auction/FunctionalCorrectness.v @@ -0,0 +1,72 @@ +Require Import backend.MachineModel. +Require Import cclib.Integers. +Require Import LayerAUCTION. +Require Import backend.Values. +Require Import backend.BuiltinSemantics. +Require Import DeepSpec.lib.Monad.StateMonadOption. +Require Import DataTypeOps. + + + +Obligation Tactic := try eexists; simpl; eauto; try congruence. +Program Definition simple_machine_env + (address origin caller callvalue + coinbase timestamp number : int256) : machine_env := + {| me_query q := + match q with + | Qcall0 Baddress => Vint address + | Qcall0 Borigin => Vint origin + | Qcall0 Bcaller => Vint caller + | Qcall0 Bcallvalue => Vint callvalue + | Qcall0 Bcoinbase => Vint coinbase + | Qcall0 Btimestamp => Vint timestamp + | Qcall0 Bnumber => Vint number + | Qcall1 Bsha_1 v => Vunit (* todo. *) + | Qcall1 Bbalance v => Vunit (* todo. *) + | Qcall1 Bblockhash v => Vunit (* todo. *) + | Qcall2 Bsha_2 (Vptr a1p) (Vint a2i) => Vptr (Ihash a1p a2i) + | Qcall2 Bsha_2 _ _ => Vunit (* todo. *) + end ; + me_transfer _ _ _ _ := True; + me_callmethod _ _ _ _ _ _ _ := True + |}. + +Definition from_option {A:Type} (default:A) (x:option A) :A := + match x with + | None => default + | Some y => y + end. + +Require Import ZArith. +Require Import cclib.Maps. + +Section WithMem. + +Import core.MemoryModel. + +Context {HmemOps: MemoryModelOps mem}. + +Section AuctionVerification. + + Parameter owner : int256. + Parameter contract_address : int256. + Parameter deadline : int256. + + Context (init_coinbase + init_timestamp + init_number : int256). + + Transparent OpenAuction_initialize_opt. + Transparent HyperTypeInst.builtin0_caller_impl HyperTypeInst.builtin0_callvalue_impl HyperTypeInst.builtin0_coinbase_impl HyperTypeInst.builtin0_timestamp_impl HyperTypeInst.builtin0_number_impl. + + Definition OpenAuction_initial := + Eval cbv -[Int256Tree.empty Int256.zero] + in + from_option init_global_abstract_data + (execStateT (OpenAuction_initialize_opt + deadline + (simple_machine_env contract_address owner owner Int256.zero + init_coinbase init_timestamp init_number)) + init_global_abstract_data). + + \ No newline at end of file diff --git a/contracts/auction/test.js b/contracts/auction/test.js deleted file mode 100755 index 0a471dc..0000000 --- a/contracts/auction/test.js +++ /dev/null @@ -1,131 +0,0 @@ -#!/usr/bin/env node - -// TODO: test withdrawals - -const fs = require('fs'); -const ethers = require('ethers'); -const assert = require('assert'); - -const bn = ethers.utils.bigNumberify; -const pe = ethers.utils.parseEther; - -if (process.argv.length != 3) { - console.log("invalid args"); - process.exit(1); -} - -const endpoint = "http://localhost:8545"; -const provider = new ethers.providers.JsonRpcProvider(endpoint); - -const abi = [ - "constructor()", - "function initialize(uint) public", - "function getbid() public view returns (uint)", - "function getbidder() public view returns (address)", - "function getchair() public view returns (address)", - "function getdeadline() public view returns (uint)", - "function bid() public returns (bool)" -]; - -const bytecode = fs.readFileSync(process.argv[2]).toString().replace(/\n|\t|\r| /g, ""); -const DEADLINE = 1000; -const CREATOR = provider.getSigner(0); - -async function deploy(deadline) { - let factory = new ethers.ContractFactory(abi, bytecode, CREATOR); - let auction = await factory.deploy(); - await auction.deployed(); - let tx = await auction.initialize(deadline); - await tx.wait(); - return auction.address; -} - -async function read(auction) { - let bid = await auction.getbid (); - let bidder = await auction.getbidder (); - let chair = await auction.getchair (); - let deadline = await auction.getdeadline (); - return [bid, bidder, chair, deadline]; -} - -async function show(auction) { - let bid, bidder, chair, deadline; - [bid, bidder, chair, deadline] = await read(auction); - console.log("bid : " + bid ); - console.log("bidder : " + bidder ); - console.log("chair : " + chair ); - console.log("deadline : " + deadline); -} - -async function testConstructor() { - let addr = await deploy(DEADLINE); - let signer = provider.getSigner(1); - let auction = new ethers.Contract(addr, abi, signer); - let bid, bidder, chair, deadline; - - [bid, bidder, chair, deadline] = await read(auction); - assert(bid.eq(bn(0))); // bid == 0 - assert(deadline.eq(bn(DEADLINE))); // deadline == DEADLINE - assert.strictEqual(bidder, await CREATOR.getAddress()); - assert.strictEqual(chair , await CREATOR.getAddress()); - console.log("constructor: passing"); -} - -async function testBids () { - let alice = provider.getSigner(1); - let bob = provider.getSigner(2); - let bid, bidder, chair, deadline; - let addr, auction, tx, overrides; - const BID0 = pe('0.1'); - const BID1 = pe('0.2'); - - // deploy - addr = await deploy(DEADLINE); - auction = new ethers.Contract(addr, abi, alice); - - // valid bid - tx = await auction.bid({ value: BID0 }); - [bid, bidder, chair, deadline] = await read(auction); - assert(bid.eq(BID0)); - assert(deadline.eq(bn (DEADLINE))); - assert.strictEqual(bidder, await alice.getAddress() ); - assert.strictEqual(chair , await CREATOR.getAddress()); - console.log("valid bid: passing"); - - // invalid bid (same bidder) - tx = await auction.bid({ value: BID1 }); - [bid, bidder, chair, deadline] = await read(auction); - assert(bid.eq(BID0)); - assert(deadline.eq(bn(DEADLINE))); - assert.strictEqual(bidder, await alice.getAddress() ); - assert.strictEqual(chair , await CREATOR.getAddress()); - console.log("invalid bid (bidder): passing"); - - // invalid bid (insufficient value) - auction = new ethers.Contract (addr, abi, bob); - tx = await auction.bid({ value: pe ('0.0') }); - [bid, bidder, chair, deadline] = await read(auction); - assert(bid.eq(BID0)); - assert(deadline.eq(bn(DEADLINE))); - assert.strictEqual(bidder, await alice.getAddress() ); - assert.strictEqual(chair , await CREATOR.getAddress()); - console.log("invalid bid (value): passing"); - - // invalid bid (deadline) - addr = await deploy(0); - auction = new ethers.Contract(addr, abi, alice); - tx = await auction.bid({ value: BID0 }); - [bid, bidder, chair, deadline] = await read(auction); - assert(bid.eq(pe('0.0'))); - assert(deadline.eq(bn(0))); - assert.strictEqual(bidder, await CREATOR.getAddress()); - assert.strictEqual(chair , await CREATOR.getAddress()); - console.log("invalid bid (deadline): passing"); -} - -async function main() { - await testConstructor(); - await testBids(); -} - -main(); diff --git a/contracts/olive/Olive.ds b/contracts/olive/Olive.ds old mode 100644 new mode 100755 index e001a13..f8f42a3 --- a/contracts/olive/Olive.ds +++ b/contracts/olive/Olive.ds @@ -133,10 +133,6 @@ layer LOCKABLE : [{ ownable : OwnableInterface }] (* we cannot have "to" which is a token that could cause parse error; need to mention in reference *) -type Wrapper = { - val: mapping[address] uint -} - object signature ERC20Interface = { constructor : unit -> unit; const balanceOf : address -> uint; @@ -151,9 +147,8 @@ object signature ERC20Interface = { object Token (pausable : PausableInterface, ownable : OwnableInterface, lockable : LockableInterface) : ERC20Interface { let _totalSupply : uint := 0u10000000000 - let wrapper : Wrapper := { val = mapping_init } let _balances : mapping[address] uint := mapping_init - let _allowed : mapping [address] Wrapper := mapping_init + let _allowed : mapping[address] mapping[address] uint := mapping_init let constructor () = let a = _totalSupply in @@ -179,7 +174,7 @@ object Token (pausable : PausableInterface, ownable : OwnableInterface, lockable true let allowance (owner, spender) = - let res = _allowed[owner].val[spender] in + let res = _allowed[owner][spender] in res let transferFrom (from_, to_, value, now) = @@ -187,7 +182,7 @@ object Token (pausable : PausableInterface, ownable : OwnableInterface, lockable assert(check <> false); pausable.whenNotPaused(); let from_bal = _balances[from_] in - let from_allow = _allowed[from_].val[msg_sender] in + let from_allow = _allowed[from_][msg_sender] in let to_bal = _balances[to_] in assert(to_ <> address(0)); assert(value <= from_bal); @@ -195,13 +190,13 @@ object Token (pausable : PausableInterface, ownable : OwnableInterface, lockable _balances[from_] := from_bal - value; (* todo: overflow check for the addition. *) _balances[to_] := to_bal + value; - _allowed[from_].val[msg_sender] := from_allow- (value); + _allowed[from_][msg_sender] := from_allow- (value); emit Transfer(from_, to_, value); true let approve (spender, value) = (* pausable.whenNotPaused(); *) - _allowed[msg_sender].val[spender] := value; + _allowed[msg_sender][spender] := value; emit Approval(msg_sender, spender, value); true @@ -209,22 +204,22 @@ object Token (pausable : PausableInterface, ownable : OwnableInterface, lockable let increaseApproval(spender, addedValue) = pausable.whenNotPaused(); (* todo: overflow check *) - let allowmsgsender = _allowed[msg_sender].val[spender] in - _allowed[msg_sender].val[spender] := allowmsgsender + addedValue; + let allowmsgsender = _allowed[msg_sender][spender] in + _allowed[msg_sender][spender] := allowmsgsender + addedValue; emit Approval(msg_sender, spender, allowmsgsender); true let decreaseApproval (_spender, _subtractedValue) = pausable.whenNotPaused(); - let oldValue = _allowed[msg_sender].val[_spender] in + let oldValue = _allowed[msg_sender][_spender] in if (_subtractedValue > oldValue) then begin - _allowed[msg_sender].val[_spender] := 0u0; - let a = _allowed[msg_sender].val[_spender] in + _allowed[msg_sender][_spender] := 0u0; + let a = _allowed[msg_sender][_spender] in emit Approval(msg_sender, _spender, a); true end else begin - _allowed[msg_sender].val[_spender] := oldValue - _subtractedValue; - let b = _allowed[msg_sender].val[_spender] in + _allowed[msg_sender][_spender] := oldValue - _subtractedValue; + let b = _allowed[msg_sender][_spender] in emit Approval(msg_sender, _spender, b); true end diff --git a/contracts/olive/test.js b/contracts/olive/test.js deleted file mode 100755 index c44364b..0000000 --- a/contracts/olive/test.js +++ /dev/null @@ -1,147 +0,0 @@ -#!/usr/bin/env node - -const fs = require('fs'); -const ethers = require("ethers"); - -if (process.argv.length != 3) { - console.log("invalid args"); - process.exit(1); -} - -const endpoint = "http://localhost:8545"; -const provider = new ethers.providers.JsonRpcProvider(endpoint); - -const abi = [ {"type":"constructor", - "name":"constructor", - "inputs":[], - "outputs":[], - "payable":false, - "constant":false, - "stateMutability":"nonpayable"}, - {"type":"function", - "name":"balanceOf", - "inputs":[{"name":"addr", "type":"address"}], - "outputs":[{"name":"", "type":"uint256"}], - "payable":false, - "constant":true, - "stateMutability":"view"}, - {"type":"function", - "name":"transfer", - "inputs":[{"name":"to_", "type":"address"},{"name":"value", "type":"uint256"},{"name":"now", "type":"uint256"}], - "outputs":[{"name":"", "type":"bool"}], - "payable":true, - "constant":false, - "stateMutability":"payable"}, - {"type":"function", - "name":"allowance", - "inputs":[{"name":"owner", "type":"address"},{"name":"spender", "type":"address"}], - "outputs":[{"name":"", "type":"uint256"}], - "payable":true, - "constant":false, - "stateMutability":"payable"}, - {"type":"function", - "name":"transferFrom", - "inputs":[{"name":"from_", "type":"address"},{"name":"to_", "type":"address"},{"name":"value", "type":"uint256"},{"name":"now", "type":"uint256"}], - "outputs":[{"name":"", "type":"bool"}], - "payable":true, - "constant":false, - "stateMutability":"payable"}, - {"type":"function", - "name":"approve", - "inputs":[{"name":"spender", "type":"address"},{"name":"value", "type":"uint256"}], - "outputs":[{"name":"", "type":"bool"}], - "payable":true, - "constant":false, - "stateMutability":"payable"}, - {"type":"function", - "name":"increaseApproval", - "inputs":[{"name":"spender", "type":"address"},{"name":"addedValue", "type":"uint256"}], - "outputs":[{"name":"", "type":"bool"}], - "payable":true, - "constant":false, - "stateMutability":"payable"}, - {"type":"function", - "name":"decreaseApproval", - "inputs":[{"name":"_spender", "type":"address"},{"name":"_subtractedValue", "type":"uint256"}], - "outputs":[{"name":"", "type":"bool"}], - "payable":true, - "constant":false, - "stateMutability":"payable"}, - {"type":"event", - "name":"OwnershipTransferred", - "inputs":[{"name":"previousOwner", "type":"address", "internalType": "address", "indexed": true}, {"name":"newOwner", "type":"address", "internalType": "address", "indexed": true}]}, - {"type":"event", - "name":"Pause", - "inputs":[{"name":"v", "type":"bool", "internalType": "bool", "indexed": false}]}, - {"type":"event", - "name":"Unpause", - "inputs":[{"name":"v", "type":"bool", "internalType": "bool", "indexed": false}]}, - {"type":"event", - "name":"Transfer", - "inputs":[{"name":"from", "type":"address", "internalType": "address", "indexed": true}, {"name":"to_", "type":"address", "internalType": "address", "indexed": true}, {"name":"value", "type":"uint256", "internalType": "uint256", "indexed": false}]}, - {"type":"event", - "name":"Approval", - "inputs":[{"name":"owner", "type":"address", "internalType": "address", "indexed": true}, {"name":"spender", "type":"address", "internalType": "address", "indexed": true}, {"name":"value", "type":"uint256", "internalType": "uint256", "indexed": false}]}] - - -const bytecode = fs.readFileSync(process.argv[2]).toString().replace(/\n|\t|\r| /g, ""); -const signer = provider.getSigner(0); -const creator = signer.getAddress(); - -async function deploy() { - console.log("sending creation transaction...") - let factory = new ethers.ContractFactory(abi, bytecode, signer); - let contract = await factory.deploy(); - await contract.deployed(); - console.log("contract address: " + contract.address); - console.log("transaction hash: " + contract.deployTransaction.hash); - let deployedBytecode = await provider.getCode(contract.address); - console.log("deployed bytecode: " + deployedBytecode); - - // contract.on("Transfer", (from, to, val) => { - // console.log("Transfer event"); - // console.log(from, to, val); - // }); - - console.log("calling balanceOf creator..."); - let balanceOf = await contract.balanceOf(creator); - console.log("value get as: " + balanceOf); - // console.log(balanceOf); - let alice = ethers.utils.getAddress("0x0000000000000000000000000000000000000024"); // arbitrary address - - // console.log("calling approve..."); - // let approve_res = await contract.approve(creator, 1000); - // console.log("approve result is " + approve_res); - - - // console.log("calling allowance"); - // let allowance_res = await contract.allowance(creator, alice); - // console.log("allowance result is " + allowance_res); - - console.log("calling approve... "); - let approve_res2 = await contract.approve(alice, 100); - console.log("approve result is " + approve_res2); - - console.log("calling allowance "); - let allowance_res2 = await contract.allowance(creator, alice); - console.log("allowance result is " + allowance_res2); - - console.log("calling transfer..."); - // to_, value, now - let transfer_res = await contract.transfer(alice, 100, 100); - console.log("transfer result is: " + transfer_res); - - - - console.log("calling balanceOf alice... after receive the token"); - let balanceOfreceiver = await contract.balanceOf(alice); - console.log("value get as after receiving: " + balanceOfreceiver); - - console.log("calling balanceOf creator... after receive the token"); - let balanceOfcreator = await contract.balanceOf(creator); - console.log("value get as after receiving: " + balanceOfcreator); - //console.log(balanceOf); - -} - -deploy(); diff --git a/contracts/token/test.js b/contracts/token/test.js deleted file mode 100755 index 1d83a5c..0000000 --- a/contracts/token/test.js +++ /dev/null @@ -1,83 +0,0 @@ -#!/usr/bin/env node - -// const net = require('net'); -// const solc = require('solc'); -const fs = require('fs'); -const ethers = require("ethers"); - -if (process.argv.length != 3) { - console.log("invalid args"); - process.exit(1); -} - -const endpoint = "http://localhost:8545"; -const provider = new ethers.providers.JsonRpcProvider(endpoint); - -// const abi = [ -// "constructor()", -// "function initialize () public returns (uint)", -// "function totalSupply() view returns (uint)", -// "function balanceOf(uint) view returns (uint)", -// "function transfer(uint, int) returns (bool)" -// ]; - -const abi = [ {"type":"function", - "name":"initialize", - "inputs":[], - "outputs":[{"name":"", "type":"uint256"}], - "payable":"true", - "constant":false, - "stateMutability":"payable"}, - {"type":"function", - "name":"totalSupply", - "inputs":[], - "outputs":[{"name":"", "type":"uint256"}], - "payable":"true", - "constant":true, - "stateMutability":"view"}, - {"type":"function", - "name":"balanceOf", - "inputs":[{"name":"tokenOwner", "type":"uint256"}], - "outputs":[{"name":"", "type":"uint256"}], - "payable":"true", - "constant":true, - "stateMutability":"view"}, - {"type":"function", - "name":"transfer", - "inputs":[{"name":"toA", "type":"uint256"},{"name":"tokens", "type":"uint256"}], - "outputs":[{"name":"", "type":"bool"}], - "payable":"true", - "constant":false, - "stateMutability":"payable"}]; - -const bytecode = fs.readFileSync(process.argv[2]).toString().replace(/\n|\t|\r| /g, ""); -const signer = provider.getSigner(0); -const creator = signer.getAddress(); - -async function deploy() { - console.log("sending creation transaction...") - let factory = new ethers.ContractFactory(abi, bytecode, signer); - let contract = await factory.deploy(); - await contract.deployed(); - console.log("contract address: " + contract.address); - console.log("transaction hash: " + contract.deployTransaction.hash); - let deployedBytecode = await provider.getCode(contract.address); - // console.log("deployed bytecode: " + deployedBytecode); - - console.log("calling initalize..."); - let tx = await contract.initialize(); - console.log("transaction hash: " + tx.hash); - - let alice = 24; // arbitrary address - console.log("calling transfer..."); - tx = await contract.transfer(alice, 100); - console.log("transaction hash: " + tx.hash); - let supply = await contract.totalSupply(); - let aliceBalance = await contract.balanceOf(alice); - let creatorBalance = await contract.balanceOf(creator); - console.log("total supply: " + supply); - console.log("creator balance: " + creatorBalance); - console.log("alice balance: " + aliceBalance); -} - -deploy(); diff --git a/contracts/token/token.ds b/contracts/token/token.ds old mode 100644 new mode 100755 diff --git a/contracts/token/token/FunctionalCorrectness.v b/contracts/token/token/FunctionalCorrectness.v new file mode 100755 index 0000000..dabf404 --- /dev/null +++ b/contracts/token/token/FunctionalCorrectness.v @@ -0,0 +1,81 @@ +Require Import token.DataTypeOps. +Require Import token.LayerFIXEDSUPPLYTOKEN. +Require Import token.Invariant. + +Require Import DeepSpec.lib.Monad.StateMonadOption. +Require Import DeepSpec.lib.Monad.RunStateTInv. +Require Import lib.ArithInv. +Import DeepSpec.lib.Monad.Monad.MonadNotation. + +Require Import ZArith. +Require Import cclib.Maps. +Require Import cclib.Integers. + +Section WithMem. + +Import core.MemoryModel. + +(* +(* Todo: move this lemma to inv_arith. *) +Lemma cmpu_Cne_true : forall x y, Int256.cmpu Cne x y = true -> x<>y. +Proof. + intros. + unfold Int256.cmpu in H. + rewrite Bool.negb_true_iff in H. + apply Int256eq_false in H. + assumption. +Qed. *) + +(* We have now loaded the specification of the transfer method. *) +(* Print FixedSupplyToken_transfer_opt. *) + +Transparent balances_sum. + +Context {memModelOps : MemoryModelOps mem}. + +(* We can now prove that the transfer method does not create or destroy tokens. *) +Theorem transfer_constant_balances_sum : forall toA n d d' me b, + runStateT (FixedSupplyToken_transfer_opt toA n me) d = Some (b, d') + -> balances_sum d' = balances_sum d. +Proof. + intros. + Transparent FixedSupplyToken_transfer_opt. + unfold FixedSupplyToken_transfer_opt in H. + inv_runStateT. + subst. + inv_arith. + unfold balances_sum. + subst. + autorewrite with updates. + remember (FixedSupplyToken_balances m2) as m. + rewrite Int256Tree_Properties.sum_swap by congruence. + apply Int256Tree_Properties.constant_sum'. + + reflexivity. + + reflexivity. + + congruence. +Qed. + +(* And similar for the transferFrom method. (The proof is actually + identical.) *) +Theorem transferFrom_constant_balances_sum : forall fromA toA n d d' me b, + runStateT (FixedSupplyToken_transferFrom_opt fromA toA n me) d = Some (b, d') + -> balances_sum d' = balances_sum d. +Proof. + intros. + Transparent FixedSupplyToken_transferFrom_opt. + unfold FixedSupplyToken_transferFrom_opt in H. + inv_runStateT. + subst. + inv_arith. + unfold balances_sum. + subst. + autorewrite with updates. + remember (FixedSupplyToken_balances m2) as m. + rewrite Int256Tree_Properties.sum_swap by congruence. + apply Int256Tree_Properties.constant_sum'. + + reflexivity. + + reflexivity. + + congruence. +Qed. + +End WithMem. diff --git a/contracts/token/token/Invariant.v b/contracts/token/token/Invariant.v new file mode 100755 index 0000000..914975f --- /dev/null +++ b/contracts/token/token/Invariant.v @@ -0,0 +1,25 @@ +Require Import ZArith. + +Require Import cclib.Maps. +Require Import token.DataTypeOps. + +(* This must match the constant defined in the token.ds file, it would be + better if the frontend generated this definition. *) +Definition _totalSupply : Z := 100000 %Z. + +Definition balances_sum (d : global_abstract_data_type) : Z := + Int256Tree_Properties.sum (FixedSupplyToken_balances d). + +Definition balances_nonnegative d := + forall i n, + Int256Tree.get i (FixedSupplyToken_balances d) = Some n + -> (0 <= n)%Z. + +(* Additionally we need that all the balance entries are non-negative, + but we will get that from the hashmap ft_cond. *) + +Definition inv d := + balances_nonnegative d /\ + balances_sum d = _totalSupply. + +Global Opaque balances_nonnegative balances_sum. (* to prevent it from being expanded when defining the VC. *) \ No newline at end of file diff --git a/contracts/token/token/ManualRefinement.v b/contracts/token/token/ManualRefinement.v new file mode 100755 index 0000000..f617932 --- /dev/null +++ b/contracts/token/token/ManualRefinement.v @@ -0,0 +1,549 @@ +Require Import backend.phase.MiniC.Language. +Require Import backend.phase.MiniC.BigstepSemantics. +Require Import backend.phase.MiniC.Semantics. +Require Import backend.Values. +Require Import backend.MemoryModel. +Require Import core.MemoryModel. + +Section STMT_FUNC. + +Context `{HM : HyperMem}. + +(* -----------------stuff that should go in core.SynthesisStmt ------------------ *) + +Require Import cclib.Maps. +Require Import core.HyperType. +Require Import DeepSpec.core.SEnv. +Require Import DeepSpec.core.Syntax. + + (* todo: move this: *) + Definition empty_temp_env := PTree.empty val. + + Let env := empty_temp_env. + + (*Variable m0 : MemLow.*) + (* --------------------- actual refinement proof ------------------ *) + +Require Import DeepSpec.Runtime. +Require Import token.EdsgerIdents. +Require Import token.DataTypes. +Require Import token.DataTypeOps. +Require Import token.DataTypeProofs. +Require Import token.LayerFIXEDSUPPLYTOKEN. +Require Import token.Invariant. +Require Import token.ObjFixedSupplyTokenCodeProofs. +Require Import token.RefineFIXEDSUPPLYTOKEN. + + +Require Import lib.Monad.StateMonad. + +Context (ge : genv) (me : MachineModel.machine_env). + +(* Note, later we should fix the the (fst ml) by making the operational + semantics thread through an abstract datatype also. *) + +(* Context`{LayerSpec : LayerSpecClass}. (* TODO: this line will go away when we fix the type of HyperMem. *) *) +Set TypeClasses Debug. + +(*Context {memModelOps : MemoryModelOps storage_env}. *) +Context`{CTXT_prf : !Layer_FIXEDSUPPLYTOKEN_Context_prf}. + +Existing Instances RefineFIXEDSUPPLYTOKEN.memModelOps RefineFIXEDSUPPLYTOKEN.CTXT_prf FIXEDSUPPLYTOKEN_hypermem GlobalLayerSpec FIXEDSUPPLYTOKEN_overlay_spec FIXEDSUPPLYTOKEN_underlay_spec. + +Require Import lib.Monad.RunStateTInv. + + +(* This should go in a generated file somewhere, + probably ObjFixedSupplyTokenCode.v. *) +Instance FixedSupplyToken_balances_var_hyperltype : + HyperLType FixedSupplyToken_balances_var. +Proof. + constructor. + constructor. + - (* ltype_get_match *) + intros j d m mm dc; split. + + apply mm. + + apply mm. + - (* ltype_set_match *) + intros f j d m fc mm dc m' disjoint_eq same_d match_indirect. + constructor. + + destruct mm as [Hrelate _]. (* relate_AbData *) + rewrite same_d in *. + destruct Hrelate. + constructor; auto. + + constructor. (* match_AbData *) + split. + * exact match_indirect. + * exact fc. +Qed. + +(* This ad-hoc lemma will probably not be needed in the generic proofs. *) +Lemma match_indirect_basic: forall m ty i cv, + is_immediate_type ty -> + cval_match_indirect m ty i cv -> + cval_match (IdentExtMap.get i m) cv. +Proof. + intros m ty i cv H mm. + inversion mm; subst; try (now inversion H). + assumption. +Qed. + + +Theorem totalSupply_correctness : +(* forall wf ge M, MakeProgram.make_globalenv GetLowDataX (M, GetLowLayer) = Errors.OK ge -> *) + forall f d d', + runStateT (FixedSupplyToken_totalSupply_opt me) d = Some (f, d') -> + forall j ml (mcond : mem_match j d ml) lg, + synth_func_cond FixedSupplyToken_totalSupply FixedSupplyToken_totalSupply_wf me d -> + exists ml' lg' g' vres, + mem_match j d' ml' /\ + ht_rel (tp:=tint_Z32) f vres /\ + eval_funcall ge me (fst ml) lg + FixedSupplyToken_totalSupply_cfun (nil (*<-args*)) + (fst ml') lg' g' vres. +Proof. + intros f d d' Hrun j ml mcond lg Hvc. + assert (mcond_copy := mcond). + destruct mcond_copy as [? [[Hmatch_balances]]]. + Transparent FixedSupplyToken_totalSupply_opt. + unfold FixedSupplyToken_totalSupply_opt in Hrun. + inv_runStateT. + subst. + + unfold FixedSupplyToken_totalSupply_cfun. + +(* We first want to build a lens corresponding to the hashtable itself, + and then to one particular dereferenced location in it. + (LChash tint_Z32 + (LCvar FixedSupplyToken_balances_var) + (ECconst_int256 tint_U (Int256.repr 0) (Int256.repr 0))) +*) + pose (l_balance := inthash_ltype_pair FixedSupplyToken_balances_var (Int256.zero)). +g pose (balance0 := l_balance.(ltype_get) d'). + + assert (ht_ft_cond balance0 /\ + cval_match_indirect (fst ml) (unpair_ty tint_Z32) + (Ihash (Iident var_FixedSupplyToken_balances_ident) Int256.zero) + (ht_cval (balance0))) + as H_balance_get. + { + assert (balance0_hyperltype : HyperLType (inthash_ltype_pair FixedSupplyToken_balances_var Int256.zero)). + { eapply inthash_ltype. + eapply FixedSupplyToken_balances_var_hyperltype. + apply thash_int_HASH_Z_Z32_hash. + } + + destruct balance0_hyperltype as [balance0_hyperltype_dir]. + specialize (balance0_hyperltype_dir eq_refl). + destruct (ltype_get_match j d' ml mcond) as [balance0_fc balance0_match]. + { simpl. exact I. } + simpl in balance0_fc, balance0_match. + split; [exact balance0_fc | exact balance0_match]. + } + destruct H_balance_get as [balance0_fc balance0_match]. + + (* Next, for the arithmetic operation we want to prove something like + exists v, + (forall m, eval_expr ge env le m (synth_expr_expr tmp e) v) /\ + ht_rel (synth_expr_spec tmp e wf se) v + *) + assert (bin_cond : @Hbinary_cond Osub tint_Z32 tint_Z32 tint_Z32 + int_Z32_sub_impl _totalSupply balance0). + { simpl. + + cbv -[Int256.modulus zeq zle zlt Z.iter Z.le Z.lt Z.gt Z.ge Z.eqb Z.leb Z.ltb Z.geb Z.gtb Z.mul Z.div Z.modulo Z.add Z.sub Z.shiftl Z.shiftr Z.lxor Z.land Z.lor Int256.add Int256.sub Int256.mul Int256.modu Int256.divu Int256.cmpu Int256.not Int256.and Int256.or Int256.xor Int256.shl Int256.shru Ziteri Z.of_nat List.length + (*ret bind mzero get put gets guard modify runStateT evalStateT execStateT *) + is_true bool_dec ZMap.get ZMap.set Int256Tree.get Int256Tree.set hlist_hd + + balances update_balances] + in Hvc. + destruct Hvc as [[_ Hvc] _]. + specialize (Hvc balance0 d'). + destruct Hvc as [_ Hvc]. + { + reflexivity. + } + apply Hvc. + auto. + } + + apply match_indirect_basic in balance0_match; [|constructor]. + + + assert (fc_totalSupply : ht_ft_cond (tp:=tint_Z32) _totalSupply). + { + split; reflexivity. + } + assert (bin_correct := Hbinary_correct _ _ + (ht_cval (tp:=tint_Z32) _totalSupply) + (ht_cval (tp:=tint_Z32) balance0) + fc_totalSupply balance0_fc + bin_cond eq_refl eq_refl). + unfold ht_cval_some in bin_correct. + assert (bin_correct' : + Some (ht_cval (tp:=tint_Z32) (Hbinary (HyperBinaryImpl:=int_Z32_sub_impl) (tpl:=tint_Z32) _totalSupply balance0)) = + (cval_binary_operation Osub (ht_cval (tp:=tint_Z32) _totalSupply) + (unpair_ty tint_Z32) (ht_cval balance0) + (unpair_ty tint_Z32))). + { remember _totalSupply as _totalSupply'. + inversion bin_correct. + simpl. + f_equal. + subst. + exact H3. + } + clear bin_correct. + symmetry in bin_correct'. + + assert (Hmatch_totalSupply : cval_match (Vint (Int256.repr _totalSupply)) (ht_cval (tp:=tint_Z32) _totalSupply)). + { + apply CVMval. + } + eapply (cval_sem_binary_operation _ _ _ _ _ _ _ _ Hmatch_totalSupply balance0_match) in bin_correct'. + destruct bin_correct' as [v [bin_correct_step bin_correct_match]]. + + + destruct (TempModel.function_initial_temps_defined 12%positive tint ftype FixedSupplyToken_totalSupply_cfun (@PTree.empty _ : temp_env)) as [v12 Hv12] ;[simpl; auto | simpl;tauto | ]. + + destruct (TempModel.function_initial_temps_defined 10%positive tint ftype FixedSupplyToken_totalSupply_cfun (@PTree.empty _ : temp_env)) as [v10 Hv10] ;[simpl; auto | simpl;tauto | ]. + + + (* Yay, that's all the setup we need, now we can run the operational + semantics. *) + + eexists. + eexists. + eexists. + eexists. + split; [|split]. + - exact mcond. + - (* The return value. *) + exact bin_correct_match. + - eapply eval_funcall_internal. + + reflexivity. + + match goal with + [|- exec_stmt _ _ ?TEMPS _ _ _ _ _ _ _ _] => + set (the_temps := TEMPS) + end. + simpl. + + eapply exec_Sseq_1. + eapply exec_Sseq_1. + eapply exec_Sset. + { exact Hv12. } + eapply eval_Elvalue. + eapply eval_Ehashderef. + eapply eval_Evar. + eapply eval_Econst_int256. + + (* eapply exec_Sseq_1. *) + eapply exec_Sset. + { exact Hv10. } + eapply eval_Ebinop. + eapply eval_Econst_int256. + eapply eval_Etempvar. rewrite PTree.gss. reflexivity. + + exact bin_correct_step. + + eapply exec_Sreturn_some. + eapply eval_Etempvar. + rewrite PTree.gss. + reflexivity. +Qed. + +(* + +Theorem transfer_correctness : +(* forall wf ge M, MakeProgram.make_globalenv GetLowDataX (M, GetLowLayer) = Errors.OK ge -> *) + forall toA toA_val tokens tokens_val f d d', + (* These things come from "good_values" in SynthesisFunc.v *) + ht_ft_cond (tp:=tint_U) toA -> ht_rel (tp:=tint_U) toA toA_val -> + ht_ft_cond (tp:=tint_Z32) tokens -> ht_rel (tp:=tint_Z32) tokens tokens_val -> + runStateT (FixedSupplyToken_transfer_opt toA tokens me) d = Some (f, d') -> + forall j ml (mcond : mem_match j d ml) lg, + synth_func_cond FixedSupplyToken_transfer FixedSupplyToken_transfer_wf toA tokens me d -> + exists ml' lg' g' vres, + mem_match j d' ml' /\ + ht_rel (tp:=tint_bool) f vres /\ + eval_funcall ge me (fst ml) lg + FixedSupplyToken_transfer_cfun (toA_val :: tokens_val :: nil (*<-args*)) + (fst ml') lg' g' vres. +Proof. + intros toA toA_val tokens tokens_val f d d' + Hft_cond_toA Hrel_toA Hft_cond_tokens Hrel_tokens + Hrun j ml mcond lg Hvc. + assert (mcond_copy := mcond). + destruct mcond_copy as [? [[Hmatch_balances]]]. + Transparent FixedSupplyToken_transfer_opt. + unfold FixedSupplyToken_transfer_opt in Hrun. + remember (Hquery0 me) as fromA. + + assert (fromA_rel : ht_cval fromA = (CVval (MachineModel.me_query me (MachineModel.Qcall0 Hbuiltin0)))). + { + subst. + apply Hbuiltin_correct. + } + (* hide the HeqfromA equation from `subst`. *) + change (fromA = Hquery0 me) with ((fun x=>x) (fromA = Hquery0 me)) in HeqfromA. + + inv_runStateT. + destruct (Integers.Int256.cmpu Cne fromA toA)%Z eqn:?. + match goal with [H: runStateT (if ?X then _ else _) _ = _ |-_] => + destruct X eqn:? end. + inv_runStateT;subst. + (* There are three goals, corresponding to each path through the if-statements. + The proof for each one is similar. + *) + { + + unfold FixedSupplyToken_transfer_cfun. + rename x into fromA. + rename x4 into d. + + pose (l_balance_fromA := inthash_ltype_pair FixedSupplyToken_balances_var fromA). + pose (balance_fromA := l_balance_fromA.(ltype_get) d). + assert (ht_ft_cond balance_fromA /\ + cval_match_indirect (fst ml) (unpair_ty tint_Z32) + (Ihash (Iident var_FixedSupplyToken_balances_ident) fromA) + (ht_cval (balance_fromA))) + as H_balance_get. + { + assert (balance_fromA_hyperltype : HyperLType (inthash_ltype_pair FixedSupplyToken_balances_var fromA)). + { eapply inthash_ltype. + eapply FixedSupplyToken_balances_var_hyperltype. + apply thash_int_HASH_Z_Z32_hash. + } + + destruct balance_fromA_hyperltype as [balance_fromA_hyperltype_dir]. + specialize (balance_fromA_hyperltype_dir eq_refl). + destruct (ltype_get_match j d ml mcond) as [balance_fromA_fc balance_fromA_match]. + { simpl. exact I. } + simpl in balance_fromA_fc, balance_fromA_match. + split; [exact balance_fromA_fc | exact balance_fromA_match]. + } + destruct H_balance_get as [balance_fromA_fc balance_fromA_match]. + + +(* destruct (TempModel.function_initial_temps_defined2 ftype FixedSupplyToken_transfer_cfun (tokens_val :: toA_val :: nil) 11%positive tint + (PTree.set 15 (Vint Int256.zero) + (PTree.set 14 + (Vint Int256.zero) + (PTree.set 13 + (Vint Int256.zero) + (PTree.set 10 + (Vint Int256.zero) + (PTree.set 11 tokens_val (PTree.set 12 toA_val (@PTree.empty val)))))))) + as [v11 Hv11]; + [simpl; auto | reflexivity | ]. *) + + destruct (TempModel.function_initial_temps_defined2 ftype FixedSupplyToken_transfer_cfun (toA_val :: tokens_val::nil) 11%positive tint + (PTree.set 15 (Vint Int256.zero) + (PTree.set 14 + (Vint Int256.zero) + (PTree.set 13 + (Vint Int256.zero) + (PTree.set 10 + (Vint Int256.zero) + (PTree.set 11 toA_val (PTree.set 12 tokens_val (@PTree.empty val)))))))) + as [v11 Hv11]; + [simpl; auto | reflexivity | ]. + + (* Fixme: this should not actually be PTree.empty, it should be init_args. *) + destruct (TempModel.function_initial_temps_defined 13%positive tint ftype FixedSupplyToken_transfer_cfun (@PTree.empty _ : temp_env)) as [v13 Hv13]; [simpl; tauto | simpl;intuition;congruence | ]. + destruct (TempModel.function_initial_temps_defined 14%positive tint ftype FixedSupplyToken_transfer_cfun (@PTree.empty _ : temp_env)) as [v14 Hv14]; [simpl; tauto | simpl;intuition;congruence | ]. + destruct (TempModel.function_initial_temps_defined 15%positive tint ftype FixedSupplyToken_transfer_cfun (@PTree.empty _ : temp_env)) as [v15 Hv15]; [simpl; tauto | simpl;intuition;congruence | ]. + + assert (Hft_cond_fromA : ht_ft_cond (tp:=tint_U) fromA). + { + split; reflexivity. + } + assert (bin_correct_ne := Hbinary_correct (op := One) _ _ + (ht_cval (tp:=tint_U) fromA) + (ht_cval (tp:=tint_U) toA) + Hft_cond_fromA Hft_cond_toA + I eq_refl eq_refl). + unfold ht_cval_some in bin_correct_ne. + assert (bin_correct_ne' : + Some (ht_cval (tp:=tint_bool) (Hbinary (HyperBinaryImpl:=int_U_ne_impl) (tpl:=tint_U) fromA toA)) = + (cval_binary_operation One (ht_cval (tp:=tint_U) fromA) + (unpair_ty tint_U) (ht_cval toA) + (unpair_ty tint_bool))). + { inversion bin_correct_ne. + simpl. + f_equal. + } + clear bin_correct_ne. + symmetry in bin_correct_ne'. + + unfold tint_U in bin_correct_ne'. + unfold addr in bin_correct_ne'. + rewrite fromA_rel in bin_correct_ne'. + + assert (Hmatch_fromA : cval_match (Vint fromA) (CVval (MachineModel.me_query me (MachineModel.Qcall0 Hbuiltin0)))). + { + rewrite <- fromA_rel. + apply CVMval. + } + + eapply (cval_sem_binary_operation _ _ _ _ _ _ _ _ Hmatch_fromA Hrel_toA) in bin_correct_ne'. + destruct bin_correct_ne' as [v_ne [bin_correct_ne_step bin_correct_ne_match]]. + + assert (bin_cond_ge : @Hbinary_cond Oge tint_Z32 tint_Z32 tint_bool + int_Z32_ge_impl balance_fromA tokens) + by exact I. + + apply match_indirect_basic in balance_fromA_match; [|constructor]. + assert (bin_correct_ge := Hbinary_correct _ _ + (ht_cval (tp:=tint_Z32) balance_fromA) + (ht_cval (tp:=tint_Z32) tokens) + balance_fromA_fc + Hft_cond_tokens + bin_cond_ge eq_refl eq_refl). + unfold ht_cval_some in bin_correct_ge. + assert (bin_correct_ge' : + Some (ht_cval (tp:=tint_bool) (Hbinary (HyperBinaryImpl:=int_Z32_ge_impl) (tpl:=tint_Z32) balance_fromA tokens )) = + (cval_binary_operation Oge + (ht_cval balance_fromA) (unpair_ty tint_Z32) + (ht_cval (tp:=tint_Z32) tokens) + (unpair_ty tint_bool))). + { + inversion bin_correct_ge. + simpl. + f_equal. + subst. + exact H3. + } + clear bin_correct_ge. + symmetry in bin_correct_ge'. + apply (cval_sem_binary_operation _ _ _ _ _ _ _ _ balance_fromA_match Hrel_tokens) in bin_correct_ge'. + destruct bin_correct_ge' as [v_ge [bin_correct_ge_step bin_correct_ge_match]]. + + pose (ne_result := Hbinary (HyperBinaryImpl:=int_U_ne_impl) fromA toA). + pose (ge_result := Hbinary (HyperBinaryImpl:=int_Z32_ge_impl) balance_fromA tokens). + + assert (bin_cond_and : @Hbinary_cond Oand tint_bool tint_bool tint_bool + int_bool_and_impl ne_result ge_result) + by exact I. + + assert (bin_correct_and := Hbinary_correct (HyperBinaryImpl0:=int_bool_and_impl) + _ _ + (ht_cval (tp:=tint_bool) ne_result) + (ht_cval (tp:=tint_bool) ge_result) + (Hbinary_returns _ _ Hft_cond_fromA Hft_cond_toA I) + (Hbinary_returns _ _ balance_fromA_fc Hft_cond_tokens bin_cond_ge) + bin_cond_and eq_refl eq_refl). + unfold ht_cval_some in bin_correct_and. + assert (bin_correct_and' : + Some (ht_cval (tp:=tint_bool) (Hbinary (HyperBinaryImpl:=int_bool_and_impl) (tpl:=tint_bool) ne_result ge_result )) = + (cval_binary_operation Oand + (ht_cval ne_result) (unpair_ty tint_bool) + (ht_cval (tp:=tint_bool) ge_result) + (unpair_ty tint_bool))). + { + inversion bin_correct_and. + simpl. + f_equal. + subst. + exact H3. + } + clear bin_correct_and. + symmetry in bin_correct_and'. + apply (cval_sem_binary_operation _ _ _ _ _ _ _ _ bin_correct_ne_match bin_correct_ge_match) in bin_correct_and'. + destruct bin_correct_and' as [v_and [bin_correct_and_step bin_correct_and_match]]. + + + eexists. + eexists. + eexists. + eexists. + split; [|split]. + + - admit. (* exact mcond *) + - (* the return value *) + unfold ht_rel. + simpl. + apply CVMval. + - eapply eval_funcall_internal. + + reflexivity. + + match goal with + [|- exec_stmt _ _ ?TEMPS _ _ _ _ _ _ _ _] => + set (the_temps := TEMPS) + end. + simpl. + + eapply exec_Sseq_1. + eapply exec_Sseq_1. + eapply exec_Sset. + { exact Hv13. } + eapply eval_Ecall0. + reflexivity. + eapply exec_Sseq_1. + eapply exec_Sset. + { rewrite PTree.gso. + exact Hv14. + congruence. + } + eapply eval_Elvalue. + eapply eval_Ehashderef. + eapply eval_Evar. + eapply eval_Etempvar. + { unfold ht_cval in fromA_rel. + simpl in fromA_rel. + inversion fromA_rel as [fromA_rel']. + symmetry. + rewrite PTree.gss. + f_equal. + exact fromA_rel'. + } + eapply exec_Sseq_1. + eapply exec_Sset. + { repeat rewrite PTree.gso by congruence. + exact Hv15. + } + + eapply eval_Elvalue. + eapply eval_Ehashderef. + eapply eval_Evar. + eapply eval_Etempvar. + { + repeat rewrite PTree.gso by congruence. + unfold the_temps. + (* stuck here exact Hv11. *) admit. + } + + (* We are in the "then" branch of the if statement. *) + eapply exec_Sifthenelse. + { eapply eval_Ebinop. (* and *) + - eapply eval_Ebinop. (* <> *) + + eapply eval_Etempvar. + rewrite !PTree.gso by congruence. + rewrite PTree.gss. + reflexivity. + + eapply eval_Etempvar. + simpl in the_temps. + unfold the_temps. + simpl. + reflexivity. + + match goal with [|- sem_binary_operation One ?X _ _ _ = _] => + replace X with (Vint fromA) + by (inversion fromA_rel as [fromA_rel_eq]; rewrite fromA_rel_eq; reflexivity) + end. + exact bin_correct_ne_step. + - eapply eval_Ebinop. (* >= *) + + eapply eval_Etempvar. + rewrite !PTree.gso by congruence. + rewrite PTree.gss. + reflexivity. + + eapply eval_Etempvar. + simpl in the_temps. + unfold the_temps. + simpl. + reflexivity. + + exact bin_correct_ge_step. + - exact bin_correct_and_step. +*) + + +End STMT_FUNC. + + diff --git a/contracts/token/token/ObjFixedSupplyTokenCodeProofs.v b/contracts/token/token/ObjFixedSupplyTokenCodeProofs.v new file mode 100644 index 0000000..93b6963 --- /dev/null +++ b/contracts/token/token/ObjFixedSupplyTokenCodeProofs.v @@ -0,0 +1,249 @@ +(* Skeleton by Edgser for token.ds *) +Require Import BinPos. +Require Import DeepSpec.Runtime. +Require Import DeepSpec.core.SynthesisFunc. +Require Import token.EdsgerIdents. +Require Import token.DataTypes. +Require Import token.DataTypeOps. +Require Import token.DataTypeProofs. +(*Require Import liblayers.compcertx.MakeProgram. +Require Import liblayers.compcertx.MemWithData.*) + +Require Import token.LayerFIXEDSUPPLYTOKEN. +Require Import lib.ArithInv. +Require Import lib.Monad.RunStateTInv. +Require Import token.Invariant. +Transparent balances_sum balances_nonnegative. + +Section EdsgerGen. + +Existing Instance GlobalLayerSpec. + +Existing Instances FIXEDSUPPLYTOKEN_overlay_spec + FIXEDSUPPLYTOKEN_data_ops + FIXEDSUPPLYTOKEN_data. + +Context {memModelOps : MemoryModelOps mem}. + +Lemma FixedSupplyToken_constructor_vc me d : + high_level_invariant d -> + synth_func_cond FixedSupplyToken_constructor FixedSupplyToken_constructor_wf + me d. +Proof. + intros. + apply FixedSupplyToken_constructor_cond_eq; auto. + unfold FixedSupplyToken_constructor_cond. + intuition. +Qed. + +Lemma FixedSupplyToken_constructor_oblg me d : + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_constructor FixedSupplyToken_constructor_wf + me d. +Proof. + intros. + apply FixedSupplyToken_constructor_obligation_eq; auto. + unfold FixedSupplyToken_constructor_obligation. + intuition. +Qed. + + +Lemma FixedSupplyToken_totalSupply_vc me d : + high_level_invariant d -> + synth_func_cond FixedSupplyToken_totalSupply FixedSupplyToken_totalSupply_wf + me d. +Proof. + intros. + apply FixedSupplyToken_totalSupply_cond_eq; auto. + unfold FixedSupplyToken_totalSupply_cond. + intuition. + inv_runStateT. + subst. + apply Zle_ge. + apply Int256Tree_Properties.sum_bound1. + + intros. + apply Zle_ge. + unfold balances_nonnegative in H1. + eauto. + + unfold balances_sum in H2. + omega. +Qed. + +Lemma FixedSupplyToken_totalSupply_oblg me d : + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_totalSupply FixedSupplyToken_totalSupply_wf + me d. +Proof. + intros. + apply FixedSupplyToken_totalSupply_obligation_eq; auto. + unfold FixedSupplyToken_totalSupply_obligation. + intuition. +Qed. + +Lemma FixedSupplyToken_balanceOf_vc a0 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + high_level_invariant d -> + synth_func_cond FixedSupplyToken_balanceOf FixedSupplyToken_balanceOf_wf + a0 me d. +Proof. + intros. + apply FixedSupplyToken_balanceOf_cond_eq; auto. + unfold FixedSupplyToken_balanceOf_cond. + intuition. +Qed. + +Lemma FixedSupplyToken_balanceOf_oblg a0 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_balanceOf FixedSupplyToken_balanceOf_wf + a0 me d. +Proof. + intros. + apply FixedSupplyToken_balanceOf_obligation_eq; auto. + unfold FixedSupplyToken_balanceOf_obligation. + intuition. +Qed. + + +Lemma FixedSupplyToken_transfer_vc a0 a1 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + high_level_invariant d -> + synth_func_cond FixedSupplyToken_transfer FixedSupplyToken_transfer_wf + a0 a1 me d. +Proof. + intros. + apply FixedSupplyToken_transfer_cond_eq; auto. + unfold FixedSupplyToken_transfer_cond. + intuition. + + + inv_runStateT. + subst. + inv_arith. + auto. + + inv_runStateT. + subst. + inv_arith. + assert (100000 < Int256.modulus) by reflexivity. + + assert (Hbound: Int256Tree_Properties.sum (FixedSupplyToken_balances g0) <= 100000). + { unfold balances_sum in H11. + omega. + } + + assert (Hnonnegative : forall k v, Int256Tree.get k (FixedSupplyToken_balances g0) = Some v -> v >= 0). + { + unfold balances_nonnegative in H6. + intros k v Hlookup. + specialize (H6 k v Hlookup). + omega. + } + + assert (l := Int256Tree_Properties.sum_bound2 H8 Hnonnegative Hbound). + apply Zle_lt_trans with (Int256Tree.get_default 0 a0 (FixedSupplyToken_balances g0) + Int256Tree.get_default 0 (MachineModel.me_caller me) (FixedSupplyToken_balances g0)). + omega. + apply Zle_lt_trans with 100000. + rewrite Zplus_comm. + exact l. + omega. +Qed. + +Lemma FixedSupplyToken_transfer_oblg a0 a1 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_transfer FixedSupplyToken_transfer_wf + a0 a1 me d. +Proof. + intros. + apply FixedSupplyToken_transfer_obligation_eq; auto. + unfold FixedSupplyToken_transfer_obligation. + intuition. +Qed. + +Lemma FixedSupplyToken_approve_vc a0 a1 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + high_level_invariant d -> + synth_func_cond FixedSupplyToken_approve FixedSupplyToken_approve_wf + a0 a1 me d. +Proof. + intros. + apply FixedSupplyToken_approve_cond_eq; auto. + unfold FixedSupplyToken_approve_cond. + intuition. +Qed. + +Lemma FixedSupplyToken_approve_oblg a0 a1 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_approve FixedSupplyToken_approve_wf + a0 a1 me d. +Proof. + intros. + apply FixedSupplyToken_approve_obligation_eq; auto. + unfold FixedSupplyToken_approve_obligation. + intuition. +Qed. + +Lemma FixedSupplyToken_transferFrom_vc a0 a1 a2 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + ht_ft_cond a2 -> ht_valid_ft_cond a2 -> + high_level_invariant d -> + synth_func_cond FixedSupplyToken_transferFrom FixedSupplyToken_transferFrom_wf + a0 a1 a2 me d. +Proof. + intros. + apply FixedSupplyToken_transferFrom_cond_eq; auto. + unfold FixedSupplyToken_transferFrom_cond. + intuition. + + inv_runStateT. + subst. + inv_arith. + auto. + + inv_runStateT. + subst. + unfold Int256.cmpu in *. + inv_arith. + clear H17. (* temp. *) + assert (100000 < Int256.modulus) by reflexivity. + + assert (Hbound: Int256Tree_Properties.sum (FixedSupplyToken_balances g0) <= 100000). + { unfold balances_sum in H15. + omega. + } + + assert (Hnonnegative : forall k v, Int256Tree.get k (FixedSupplyToken_balances g0) = Some v -> v >= 0). + { + unfold balances_nonnegative in H10. + intros k v Hlookup. + specialize (H10 k v Hlookup). + omega. + } + assert (l := Int256Tree_Properties.sum_bound2 H12 Hnonnegative Hbound). + + + apply Zle_lt_trans with (Int256Tree.get_default 0 a0 (FixedSupplyToken_balances g0) + Int256Tree.get_default 0 a1 (FixedSupplyToken_balances g0)). + omega. + apply Zle_lt_trans with 100000. + exact l. + omega. +Qed. + +Lemma FixedSupplyToken_transferFrom_oblg a0 a1 a2 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + ht_ft_cond a2 -> ht_valid_ft_cond a2 -> + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_transferFrom FixedSupplyToken_transferFrom_wf + a0 a1 a2 me d. +Proof. + intros. + apply FixedSupplyToken_transferFrom_obligation_eq; auto. + unfold FixedSupplyToken_transferFrom_obligation. + intuition. +Qed. + +End EdsgerGen. diff --git a/contracts/token/token/RefineFIXEDSUPPLYTOKEN.v b/contracts/token/token/RefineFIXEDSUPPLYTOKEN.v new file mode 100755 index 0000000..bf7e8f4 --- /dev/null +++ b/contracts/token/token/RefineFIXEDSUPPLYTOKEN.v @@ -0,0 +1,56 @@ +(* This should be automatically generated, but I put in a manually written version for now. *) + +Require Import BinNums. +Require Import DeepSpec.core.HyperTypeInst. +Require Import DeepSpec.Runtime. +Require Import DeepSpec.Linking. +Require Import token.EdsgerIdents. +Require Import token.DataTypes. +Require Import token.DataTypeOps. +Require Import token.DataTypeProofs. + +Require Import token.LayerFIXEDSUPPLYTOKEN. + + +(* +Context {mem}`{Hmem: Mem.MemoryModel mem}. +Context`{Hmwd: UseMemWithData mem}. +Context`{make_program_ops: !MakeProgramOps Clight.function Ctypes.type Clight.fundef Ctypes.type}. +Context`{Hmake_program: !MakeProgram Clight.function Ctypes.type Clight.fundef Ctypes.type}. +Context`{global_abdata : !GlobalAbData init_global_abstract_data global_low_level_invariant}. + +Context {memModelOps : MemoryModelOps mem}. + *) + +Existing Instances GlobalLayerSpec FIXEDSUPPLYTOKEN_overlay_spec FIXEDSUPPLYTOKEN_underlay_spec. + +Context`{CTXT_prf : !Layer_FIXEDSUPPLYTOKEN_Context_prf}. + +(* +Record relate_RData (j : meminj) (habd : GetHighData) (labd : GetLowData) : Prop + := mkrelate_RData { +}.*) + +Record match_RData (habd : GetHighData) (m : mem) (j : MemoryModel.meminj) : Prop + := MATCH_RDATA { + balances_ma : variable_match FixedSupplyToken_balances_var habd m +}. + +Local Hint Resolve MATCH_RDATA. + +Global Instance rel_ops: CompatRelOps GetHighDataX GetLowDataX := +{ + relate_AbData f d1 d2 := True; + match_AbData d1 m f := match_RData d1 m f; + new_glbl := var_FixedSupplyToken_balances_ident :: nil +}. + + +Global Instance rel_prf: CompatRel GetHighDataX GetLowDataX. + + +Existing Instance GlobalLayerSpec. + +Instance FIXEDSUPPLYTOKEN_hypermem : MemoryModel.HyperMem + := { Hcompatrel := {| crel_prf := rel_prf |} }. + diff --git a/contracts/token/token/_CoqProject b/contracts/token/token/_CoqProject new file mode 100755 index 0000000..9ddcc27 --- /dev/null +++ b/contracts/token/token/_CoqProject @@ -0,0 +1,11 @@ +-R ../../.. DeepSpec +-R . token +EdsgerIdents.v +DataTypeOps.v +DataTypeProofs.v +DataTypes.v +LayerFIXEDSUPPLYTOKEN.v +RefineFIXEDSUPPLYTOKEN.v +ObjFixedSupplyTokenCodeProofs.v +FunctionalCorrectness.v +Invariant.v \ No newline at end of file diff --git a/contracts/token/token/prf.v b/contracts/token/token/prf.v new file mode 100644 index 0000000..4b09193 --- /dev/null +++ b/contracts/token/token/prf.v @@ -0,0 +1,281 @@ +Require Import token.DataTypes. +Require Import token.DataTypeOps. +Require Import lib.Monad.StateMonadOption. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import ZArith. +Require Import core.HyperTypeInst. +Require Import backend.MachineModel. +Require Import token.LayerFIXEDSUPPLYTOKEN. + +Definition state := global_abstract_data_type. + +Definition init_state := init_global_abstract_data. + +Definition wei := Z. + +Definition addr := int256. + +Definition blocknumber := int256. + +Existing Instance GlobalLayerSpec. + +Section step. + + Context (contract_address : addr). + + Section mstep. + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Definition make_machine_env (caller: addr) + : machine_env state + := {| me_address := contract_address; + me_origin := caller; + me_caller := caller; (* need update after every control-flow transfer *) + me_callvalue := Int256.repr (0); + me_coinbase := coinbase; + me_timestamp := timestamp; + me_number := number; + me_balance := balance; + me_blockhash := blockhash; + (* not implemented *) + me_transfer _ _ _ _ _ := False; + me_callmethod _ _ _ _ _ _ _ _ _ _ := False; + me_log _ _ _ := prev_contract_state; + |}. + + Import MonadNotation. + + Definition lif {A:Type} + (caller : addr) + (cmd : machine_env state -> osT global_abstract_data_type A) + : osT state A := + st <- get;; + let me := make_machine_env caller in + match runStateT (cmd me) st with + | None => mzero + | Some (v, st') => put st' ;; ret v + end. + + (* osT state int256 = + state transformer with option monad, state is state, and value is int256 *) + Print osT. (* = fun D : Type => stateT D option *) + Print stateT. (* (S : Type) (m : Type -> Type) (t : Type) : Type := mkStateT + { runStateT : S -> m (t * S)%type } *) + + Print runStateT. (* takes monad transformer, state, returns m (t * S) *) + + (* How the state of the system changes in a single method call made by player p. *) + Inductive mstep u (st st' : state) : Prop := + | FixedSupplyToken_constructor_step : forall r , runStateT (FixedSupplyToken_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_totalSupply_step : forall r , runStateT (FixedSupplyToken_totalSupply_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_balanceOf_step : forall r tokenOwner, runStateT (FixedSupplyToken_balanceOf_opt tokenOwner (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_transfer_step : forall r toA tokens, runStateT (FixedSupplyToken_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_approve_step : forall r spender tokens, runStateT (FixedSupplyToken_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_transferFrom_step : forall r fromA toA tokens, runStateT (FixedSupplyToken_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' +. + + (* We can compute a new block by letting the players call the contract + in some arbitrary order. *) + Inductive multi_mstep : state -> state -> Prop := + | multi_mstep_reflexive : forall (st : state), multi_mstep st st + | multi_mstep_transitive : forall (st st' st'' : state) u, + multi_mstep st st' -> mstep u st' st'' -> multi_mstep st st''. + + + (* A block is sufficiently synchronous if every player got a chance to submit a + transaction to to. *) + (* TODO, I think this definition is fine for now, but it seems little too clever, + should probably re-state this using some straightforward tracking of the states + we pass through. *) + Definition multi_mstep_synchronous st1 st2 := + forall u, exists st st', + multi_mstep st1 st /\ mstep u st st' /\ multi_mstep st' st2. + + (* Here are a bunch of induction principles inspired by linear temporal logic. *) + + (* Prove that some property P holds "globally", i.e. for each state along a + path. + You can also specify a property Pprev which is known to hold for the + prev_contract_state. If not needed, it can just be True. + *) + Lemma multi_mstep_Global_ind : forall (Pprev P : state -> Prop), + (forall u st st', Pprev prev_contract_state -> P st -> mstep u st st' -> P st') -> + Pprev prev_contract_state -> forall st st', P st -> multi_mstep st st' -> P st'. + Proof. + induction 4; eauto. + Qed. + + (* + (* Prove that P holds "until" Q along a path. + "Until" is a liveness assertion, so we need the synchronicity assumption. *) + Lemma multi_mstep_Until_ind : forall (Pprev P Q : state -> Prop), + (forall p st st', Pprev prev_contract_state -> P st -> In p players + -> mstep p st st' -> (P st' \/ Q st')) -> + Pprev prev_contract_state -> + forall st, + P st -> exists st', multi_mstep st st' -> (P st' \/ Q st'). + Proof. + induction 4; eauto *) + + End mstep. + + + Definition Int256_incr x := Int256.add x Int256.one. + + Inductive bstep (n : blocknumber) : state -> state -> Prop := + | bstep_step : forall coinbase timestamp balance blockhash st st', + multi_mstep coinbase timestamp n balance blockhash st st st' -> + bstep n st st'. + + Inductive multi_bstep : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_reflexive : forall n (st : state), multi_bstep n st n st + | multi_bstep_transitive : forall n n' (st st' st'' : state), + multi_bstep n st n' st' -> bstep (Int256_incr n') st' st'' -> multi_bstep n st (Int256_incr n') st''. + + (* multi_bstep is the step relation without any synchronicity assumption. + This is sufficient to prove some safety properties, but for most interesting + theorems we instead need to use this synchronous version: *) + + Inductive bstep_synch (n : blocknumber) : state -> state -> Prop := + | bstep_synch_step : forall coinbase timestamp balance blockhash st st', + multi_mstep_synchronous coinbase timestamp n balance blockhash st st st' -> + bstep_synch n st st'. + + Inductive multi_bstep_synch : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_synch_reflexive : forall n (st : state), multi_bstep_synch n st n st + | multi_bstep_synch_transitive : forall n n' (st st' st'' : state), + multi_bstep_synch n st n' st' -> bstep_synch (Int256_incr n') st' st'' -> multi_bstep_synch n st (Int256_incr n') st''. + + Lemma multi_bstep_Global_ind : forall (P : state -> Prop), + (forall coinbase timestamp number balance blockhash prev_block u st st', + P prev_block + -> P st + -> mstep coinbase timestamp number balance blockhash prev_block u st st' + -> P st') + -> forall n st n' st', + P st -> multi_bstep n st n' st' -> P st'. + Proof. + induction 3. + - auto. + - inversion H2; subst. + eapply multi_mstep_Global_ind with (st:=st') (prev_contract_state := st'). + + intros. + refine (H _ _ _ _ _ _ _ _ _ _ _ H6); auto. + + apply IHmulti_bstep; assumption. + + apply IHmulti_bstep; assumption. + + exact H3. + Qed. + End step. + + Section DeepSEAGenericProof. + + Lemma Int256Tree_reduce : forall (i: int256) (v: Z) (t: Int256Tree.t Z), Int256Tree.get_default 0%Z i (Int256Tree.set i v t) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gss . + reflexivity. + Qed. + + Lemma Int256Tree_mreduce : forall (i j : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set i v t)) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gss. + reflexivity. + exact H. + Qed. + + Lemma Int256Tree_mireduce : forall (i j k : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + i <> k -> + j <> k -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set k v t)) = + Int256Tree.get_default 0%Z i t. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gso. + reflexivity. + exact H0. + exact H. + Qed. + + Lemma add_sub_inv : forall (i j : Z32), (i + j - i)%Z = j. + Proof. + intros. + omega. + Qed. + End DeepSEAGenericProof. + + Section Proof. + Context (* (strategies : strategy_tuple) *) + (initial_balances : addr -> Z) + (contract_address : int256). + + Context (init_bt init_rt : int256) + (init_coinbase : int256) + (init_timestamp : int256) + (init_number : int256) + (init_blockhash : int256 -> int256) + (pre_init_state init_state : state). + + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Require Import lib.Monad.RunStateTInv. + Require Import lib.ArithInv. + + Definition make_machine_env_wrapped prev_st user := + make_machine_env contract_address coinbase timestamp number balance blockhash prev_st user. + + Lemma make_machine_env_caller_eq : forall st caller, me_caller (make_machine_env_wrapped st caller) = caller. + Proof. auto. Qed. + + Lemma make_machine_env_address_eq : forall st caller, me_address (make_machine_env_wrapped st caller) = contract_address. + Proof. auto. Qed. + +Transparent FixedSupplyToken_constructor_opt. +Transparent FixedSupplyToken_totalSupply_opt. +Transparent FixedSupplyToken_balanceOf_opt. +Transparent FixedSupplyToken_transfer_opt. +Transparent FixedSupplyToken_approve_opt. +Transparent FixedSupplyToken_transferFrom_opt. + +Ltac rds := +unfold FixedSupplyToken_constructor_opt in *; +unfold FixedSupplyToken_totalSupply_opt in *; +unfold FixedSupplyToken_balanceOf_opt in *; +unfold FixedSupplyToken_transfer_opt in *; +unfold FixedSupplyToken_approve_opt in *; +unfold FixedSupplyToken_transferFrom_opt in *; +inv_runStateT; +subst; +inv_arith; +simpl; +try rewrite make_machine_env_caller_eq in *; +try rewrite make_machine_env_address_eq in *; +try rewrite Int256Tree_reduce in *; +try rewrite Int256Tree_mreduce in *; +try rewrite Int256Tree_mireduce in *; +auto. + +Theorem sample : forall n, n = n. + +End Proof. diff --git a/contracts/token_ant/token.ds b/contracts/token_ant/token_ant.ds similarity index 100% rename from contracts/token_ant/token.ds rename to contracts/token_ant/token_ant.ds diff --git a/minic_unittests/basic_constructor.mc b/minic_unittests/basic_constructor.mc new file mode 100644 index 0000000..8a7298c --- /dev/null +++ b/minic_unittests/basic_constructor.mc @@ -0,0 +1,10 @@ +int256 foo; + +void constructor() { + int256 bar; + int256 baz; + + bar = 0; + baz = bar; + foo = baz; +} diff --git a/minic_unittests/comments.mc b/minic_unittests/comments.mc new file mode 100644 index 0000000..ec0510a --- /dev/null +++ b/minic_unittests/comments.mc @@ -0,0 +1,20 @@ +void foo() { + + // single-line comment + + /* simple + multiline + comment + */ + + /* nested + /* multi /* line */ + com- */ + ments */ + + int8 /* interrupting comment */ tmp; + + // Multiline comment symbol '/*' ignored + + +} diff --git a/minic_unittests/contracts/amm/amm.ds b/minic_unittests/contracts/amm/amm.ds new file mode 100644 index 0000000..4b949ae --- /dev/null +++ b/minic_unittests/contracts/amm/amm.ds @@ -0,0 +1,386 @@ +(* TODO: add emit Events *) + +(* PROOFS: + 1. change re-defined layers + 2. balance tracks reserve + 3. increasing k -> algebra + 4. no path independence, splitting trades is strictly more expensive -> algebra, appendix D + 5. liquidity token economics is right, i.e., UNI-V2 tracks the actual share of pool --> actually, where is the trading of UNI-V2 Pair tokens? What is its value? + 6. MOD: slippage control + 7. ... cost of manipulation (hard) +*) + +object signature ERC20Interface = { + constructor : unit -> unit; + mint : address * int -> unit; + burn : address * int -> unit; + totalSupply : unit -> int; + balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object LiquidityToken : ERC20Interface { + let _totalSupply : int := 0 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + balances[msg_sender] := 100000 + + let mint (toA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply + value; + let to_bal = balances[toA] in + balances[toA] := to_bal + value + + let burn (fromA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply - value; + let from_bal = balances[fromA] in + balances[fromA] := from_bal - value + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +object signature FixedERC20Interface = { + constructor : unit -> unit; + const totalSupply : unit -> int; + balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object FixedSupplyToken : FixedERC20Interface { + let _totalSupply : int := 100000 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + _totalSupply := 100000; + balances[msg_sender] := 100000 + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +(* this AMM supports ERC20 tokens, ETH can be incorporated as WETH for trading *) +object signature AMMInterface = { + constructor : unit -> unit; + (* swap : int * int * address -> unit; *) + + coarsed simpleSwap0 : address -> int; (* transfer token0 to amm, results number of token1 swapped *) + + (* router *) + (* swapTokensForExactTokens + swapExactTokensForTokens + addLiquidity (* why restrain the liquidity to not depreciating an asset too much? *) + removeLiquidity *) + + (* fund management functions *) + coarsed mint : address -> unit; (* provide liquidity to this pair *) + coarsed burn : address -> unit; (* remove liquidity from this pair *) + + coarsed skim : address -> unit; (* arbitrage to force balances to match reserves *) + coarsed sync : unit -> unit; (* force reserves to match balances *) + k : unit -> int; (* get the constant product *) + + (* oracles *) + (* we do not provide Uniswap V2 version of smoothed oracles at this moment *) + quote0 : int -> int; (* given some amount of an asset and pair reserves, returns an equivalent amount of the other asset *) + getAmountOut0 : int -> int; (* given an input amount of an asset and pair reserves, returns the maximum output amount of the other asset *) + getAmountIn0 : int -> int; (* given an output amount of an asset and pair reserves, returns a required input amount of the other asset *) +} + +(* assuming 0.3% txn fee, excluding the 0.05% additional fee in Uniswap V2 design *) +(* we do not support liquidity tokens at the moment *) +object AutomatedMarketMaker + ( liquidityToken: ERC20Interface, + erc20Token0: FixedERC20Interface, + erc20Token1: FixedERC20Interface) : AMMInterface { + + let _token0 : address := address(0xdac17f958d2ee523a2206206994597c13d831ec7) (* USDT *) + let _token1 : address := address(0x6b175474e89094c44da98b954eedeac495271d0f) (* DAI *) + let _owner : address := address(0x0) + let _reserve0 : int := 0 + let _reserve1 : int := 0 + let blockTimestampLast : int := 0 (* new oracle *) + let price0CumulativeLast : int := 0 (* new oracle *) + let price1CumulativeLast : int := 0 (* new oracle *) + let kLast : int := 0 (* reserve0 * reserve1, as of immediately after the most recent liquidity event *) + + let constructor () = + _owner := msg_sender + + (* transfer token0 and token1 to this contract, then mint liquidity tokens *) + let mint (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0 = balance0 - reserve0 in + let amount1 = balance1 - reserve1 in + (* update reserve 0 and reserve 1 *) + let totalSupply = liquidityToken.totalSupply() in + let liquidity = if totalSupply = 0 + then + begin + liquidityToken.mint(address(0x0), 1000); (* lock the first liquidity tokens*) + amount0 * amount1 - 1000 (* in Uniswap this should be square-rooted *) + end + else + let x = amount0 * totalSupply / reserve0 in + let y = amount1 * totalSupply / reserve1 in + if x < y then x else y + in + assert (liquidity > 0); + liquidityToken.mint(toA, liquidity); + _reserve0 := balance0; + _reserve1 := balance1 + + (* need toA first transfer liquidity tokens to this contract, then operate *) + (* it does not use ERC20 style allowances, just direct transfer *) + let burn (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let liquidity = liquidityToken.balanceOf(this_address) in + let totalSupply = liquidityToken.totalSupply() in + let amount0 = liquidity * balance0 / totalSupply in + let amount1 = liquidity * balance1 / totalSupply in + assert (amount0 > 0 /\ amount1 > 0); + liquidityToken.burn(this_address, liquidity); + let success = erc20Token0.transfer(toA, amount0) in + assert (success); + let success = erc20Token1.transfer(toA, amount1) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + (* adhere to the stanford paper formalization, targeted for vanilla proof *) + let simpleSwap0 (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = balance0 - reserve0 in + let token0 = _token0 in + let token1 = _token1 in + assert (toA <> token0 /\ toA <> token1); + assert (amount0In > 0); + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let result = numerator / denominator in + let success = erc20Token1.transfer(toA, result) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1; + let resultU = result in + resultU + + (* derives from Uniswap V2 flashswap, but not flashswap since DeepSEA does not allow control-flow aggregation *) + (* first transfer tokens to this contract, then call swap to swap them *) + (* notice this does not refund, so possibly not optimal rate *) + (* let swap (amount0Out, amount1Out, toA) = + assert (amount0Out > 0 \/ amount1Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let token0 = _token0 in + let token1 = _token1 in + assert (amount0Out < reserve0 /\ amount1Out < reserve1); + assert (toA <> token0 /\ toA <> token1); + (* check the assets satisfy what is wanted *) + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + (* we can prove that balance > reserve here *) + let amount0In = helper.getAmountIn(balance0, 0, reserve0) in (* balance0 - reserve0 *) + let amount1In = helper.getAmountIn(balance1, 0, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0 - amount0Out, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1 - amount1Out, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); + + (* sending money *) + let _ = if amount0Out > 0 then + let success = erc20Token0.transfer(toA, amount0Out) in + assert (success); + () + else + () + in + let _ = if amount1Out > 0 then + let success = erc20Token1.transfer(toA, amount1Out) in + assert (success); + () + else + () + in + (* flash swap style: sent money (amount0Out, amount1Out) to msg_sender, now wait to get back *) + (* after got money back *) + (* let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = helper.getAmountIn(balance0, amount0Out, reserve0) in + let amount1In = helper.getAmountIn(balance1, amount1Out, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); *) + (* update the reserves, also update the oracle if possible *) + _reserve0 := balance0; + _reserve1 := balance1 *) + + let skim (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let skim0 = balance0 - reserve0 in + let skim1 = balance1 - reserve1 in + let success = erc20Token0.transfer(toA, skim0) in + assert (success); + let success = erc20Token1.transfer(toA, skim1) in + assert (success) + + let sync () = + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + let k () = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let resultU = reserve0 * reserve1 in + resultU + + (* Uniswap did not use Q112 library to do this computation *) + + (* given amount0 of token0, the equivalent value in token1 *) + let quote0 (amount0) = + assert (amount0 > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); (* prove that this can be dropped? *) + let resultU = amount0 * reserve1 / reserve0 in + resultU + + let getAmountOut0 (amount0In) = + assert (amount0In > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let resultU = numerator / denominator in + resultU + + let getAmountIn0 (amount0Out) = + assert (amount0Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let numerator = reserve1 * amount0Out * 1000 in + let denominator = (reserve0 - amount0Out) * 997 in + let resultU = (numerator / denominator) + 1 in + resultU + +} + +layer signature AMMLibSig = { + erc20Token0 : FixedERC20Interface; + erc20Token1 : FixedERC20Interface; + liquidityToken : ERC20Interface; +} + +layer AMMLIB : [{}] AMMLibSig = { + erc20Token0 = address(0x10033) <: FixedSupplyToken; + erc20Token1 = address(0x10032) <: (clone FixedSupplyToken); (* cloned an object *) + liquidityToken = address(0x10031) <: LiquidityToken; +} + +layer signature AMMSig = { + amm : AMMInterface +} + +layer AMM : [AMMLibSig] AMMSig = { + amm = address(0x10030) <: AutomatedMarketMaker +} + +layer COMPLETE = AMM @ AMMLIB \ No newline at end of file diff --git a/minic_unittests/contracts/amm/amm/RefineAMM.v b/minic_unittests/contracts/amm/amm/RefineAMM.v new file mode 100644 index 0000000..4137973 --- /dev/null +++ b/minic_unittests/contracts/amm/amm/RefineAMM.v @@ -0,0 +1,259 @@ +(* WARNING: This file is generated by Edsger, the DeepSEA compiler. + All modification will be lost when regenerating. *) +(* Module amm.RefineAMM for amm.ds *) +Require Import BinPos. +Require Import DeepSpec.Runtime. +Require Import DeepSpec.Linking. +Require Import amm.EdsgerIdents. +Require Import amm.DataTypes. +Require Import amm.DataTypeOps. +Require Import amm.DataTypeProofs. +Require Import layerlib.LayerCalculusLemma. +Require Import layerlib.RefinementTactic. +Require Import layerlib.RefinementTacticAux. +Require Import liblayers.compcertx.MakeProgram. +Require Import liblayers.compcertx.MemWithData. + +Require Import amm.LayerAMM. +Require Import amm.LSimAMMLIB. + +Section EdsgerGen. + + +Context {mem}`{Hmem: Mem.MemoryModel mem}. +Context`{Hmwd: UseMemWithData mem}. +Context`{make_program_ops: !MakeProgramOps Clight.function Ctypes.type Clight.fundef Ctypes.type}. +Context`{Hmake_program: !MakeProgram Clight.function Ctypes.type Clight.fundef Ctypes.type}. +Instance GlobalLayerSpec : LayerSpecClass := { + make_program_ops := make_program_ops; + Hmake_program := Hmake_program; + GetHighData := global_abstract_data_type +}. +Context`{global_abdata : !GlobalAbData init_global_abstract_data global_low_level_invariant}. +Context`{CTXT_prf : !Layer_AMM_Context_prf}. +Context`{AMMLIB_pres_inv : !AMMLIB_preserves_invariants}. +Context`{AMM_pres_inv : !AMM_preserves_invariants}. + +Existing Instances AMM_overlay_spec AMM_underlay_spec. + +Record relate_RData (j : meminj) (habd : GetHighData) (labd : GetLowData) : Prop + := mkrelate_RData { + (* FixedSupplyToken *) + FixedSupplyToken__totalSupply_re : ; + FixedSupplyToken_balances_re : ; + FixedSupplyToken_allowances_re : ; + (* FixedSupplyToken1 *) + FixedSupplyToken1__totalSupply_re : ; + FixedSupplyToken1_balances_re : ; + FixedSupplyToken1_allowances_re : ; + (* LiquidityToken *) + LiquidityToken__totalSupply_re : ; + LiquidityToken_balances_re : ; + LiquidityToken_allowances_re : +}. + +Record match_RData (habd : GetHighData) (m : mem) (j : meminj) : Prop + := MATCH_RDATA { + AutomatedMarketMaker__token0_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker__token0_var habd m; + AutomatedMarketMaker__token1_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker__token1_var habd m; + AutomatedMarketMaker__owner_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker__owner_var habd m; + AutomatedMarketMaker__reserve0_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker__reserve0_var habd m; + AutomatedMarketMaker__reserve1_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker__reserve1_var habd m; + AutomatedMarketMaker_blockTimestampLast_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker_blockTimestampLast_var habd m; + AutomatedMarketMaker_price0CumulativeLast_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker_price0CumulativeLast_var habd m; + AutomatedMarketMaker_price1CumulativeLast_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker_price1CumulativeLast_var habd m; + AutomatedMarketMaker_kLast_ma : variable_match AutomatedMarketMaker_AutomatedMarketMaker_kLast_var habd m +}. + +Local Hint Resolve MATCH_RDATA. + +Global Instance rel_ops: CompatRelOps GetHighDataX GetLowDataX := +{ + relate_AbData f d1 d2 := relate_RData f d1 d2; + match_AbData d1 m f := match_RData d1 m f; + new_glbl := var_AutomatedMarketMaker_AutomatedMarketMaker__token0_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker__token1_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker__owner_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker__reserve0_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker__reserve1_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker_blockTimestampLast_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker_price0CumulativeLast_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker_price1CumulativeLast_ident :: var_AutomatedMarketMaker_AutomatedMarketMaker_kLast_ident :: nil +}. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker__token0_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker__token0_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker__token0_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker__token1_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker__token1_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker__token1_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker__owner_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker__owner_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker__owner_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker__reserve0_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker__reserve0_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker__reserve0_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker__reserve1_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker__reserve1_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker__reserve1_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker_blockTimestampLast_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker_blockTimestampLast_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker_blockTimestampLast_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker_price0CumulativeLast_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker_price0CumulativeLast_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker_price0CumulativeLast_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker_price1CumulativeLast_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker_price1CumulativeLast_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker_price1CumulativeLast_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Global Instance AutomatedMarketMaker_AutomatedMarketMaker_kLast_hyper_ltype_static : + HyperLTypeStatic AutomatedMarketMaker_AutomatedMarketMaker_kLast_var new_glbl. +Proof. + split; try solve + [ Decision.decision + | simpl; auto + | simpl AutomatedMarketMaker_AutomatedMarketMaker_kLast_var.(ltype_offset); + rewrite Int256.unsigned_zero; apply Z.divide_0_r ]. +Qed. + +Lemma relate_incr: + forall abd abd' f f', + relate_RData f abd abd' -> + inject_incr f f' -> + relate_RData f' abd abd'. +Proof. + inversion 1; subst; intros; simpl in *. + repeat match goal with + | H : _ /\ _ |- _ => destruct H + end. + repeat (constructor; simpl; eauto). +Qed. +Global Instance rel_prf: CompatRel GetHighDataX GetLowDataX. +Proof. + constructor; [ destruct 1 .. |]; intros. + - constructor; eapply inject_match_correct; eauto with typeclass_instances. + - constructor; eapply store_match_correct; eauto with typeclass_instances. + - constructor; eapply alloc_match_correct; eauto with typeclass_instances. + - constructor; eapply free_match_correct; eauto with typeclass_instances. + - constructor; eapply storebytes_match_correct; eauto with typeclass_instances. + - eapply relate_incr; eauto. +Qed. + +(* +Local Instance: ExternalCallsOps (mwd GetLowDataX) := CompatExternalCalls.compatlayer_extcall_ops AMMLIB_Layer. +Local Instance: CompilerConfigOps _ := CompatExternalCalls.compatlayer_compiler_config_ops AMMLIB_Layer. +*) + +Instance AMM_hypermem : MemoryModel.HyperMem + := { Hcompatrel := {| crel_prf := rel_prf |} }. + +Class AMM_Underlay_preserves_invariants := { + AMM_Underlay_FixedSupplyToken_constructor_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken_constructor_spec | 5; + AMM_Underlay_FixedSupplyToken_balanceOf_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken_balanceOf_spec | 5; + AMM_Underlay_FixedSupplyToken_transfer_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken_transfer_spec | 5; + AMM_Underlay_FixedSupplyToken_approve_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken_approve_spec | 5; + AMM_Underlay_FixedSupplyToken_transferFrom_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken_transferFrom_spec | 5; + AMM_Underlay_FixedSupplyToken1_constructor_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken1_constructor_spec | 5; + AMM_Underlay_FixedSupplyToken1_balanceOf_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken1_balanceOf_spec | 5; + AMM_Underlay_FixedSupplyToken1_transfer_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken1_transfer_spec | 5; + AMM_Underlay_FixedSupplyToken1_approve_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken1_approve_spec | 5; + AMM_Underlay_FixedSupplyToken1_transferFrom_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) FixedSupplyToken1_transferFrom_spec | 5; + AMM_Underlay_LiquidityToken_constructor_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_constructor_spec | 5; + AMM_Underlay_LiquidityToken_mint_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_mint_spec | 5; + AMM_Underlay_LiquidityToken_burn_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_burn_spec | 5; + AMM_Underlay_LiquidityToken_totalSupply_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_totalSupply_spec | 5; + AMM_Underlay_LiquidityToken_balanceOf_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_balanceOf_spec | 5; + AMM_Underlay_LiquidityToken_transfer_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_transfer_spec | 5; + AMM_Underlay_LiquidityToken_approve_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_approve_spec | 5; + AMM_Underlay_LiquidityToken_transferFrom_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) LiquidityToken_transferFrom_spec | 5; + AMM_Underlay_AutomatedMarketMakerLib_constructor_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) AutomatedMarketMakerLib_constructor_spec | 5; + AMM_Underlay_AutomatedMarketMakerLib_getAmountIn_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) AutomatedMarketMakerLib_getAmountIn_spec | 5; + AMM_Underlay_AutomatedMarketMakerLib_getBalanceAdjusted_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) AutomatedMarketMakerLib_getBalanceAdjusted_spec | 5; + AMM_Underlay_AutomatedMarketMakerLib_min_preserves_invariants :> + CompatGenSem.PreservesInvariants (HD := cdataLow) AutomatedMarketMakerLib_min_spec | 5 +}. +Instance AMM'AMMLIB_preserves_invariants : AMM_Underlay_preserves_invariants. +Proof. esplit; apply AMMLIB_pres_inv. Defined. + +(* +Lemma passthrough_correct: + sim (crel (CompatRel0 := rel_prf) _ _) AMM_Layer_passthrough AMMLIB_Layer. +Proof. + Local Opaque simRR mapsto layer_mapsto_primitive. + unfold GlobalLayerSpec, MemoryModel.GetHighDataX. + simpl. + + sim_oplus; simpl. + + Local Transparent simRR mapsto layer_mapsto_primitive. +Qed.*) +End EdsgerGen. diff --git a/minic_unittests/contracts/amm/amm/prf.v b/minic_unittests/contracts/amm/amm/prf.v new file mode 100644 index 0000000..aad3bf6 --- /dev/null +++ b/minic_unittests/contracts/amm/amm/prf.v @@ -0,0 +1,436 @@ +Require Import amm.DataTypes. +Require Import amm.DataTypeOps. +Require Import lib.Monad.StateMonadOption. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import ZArith. +Require Import core.HyperTypeInst. +Require Import backend.MachineModel. +Require Import amm.LayerAMMLIB. +Require Import amm.LayerAMM. + +Definition state := global_abstract_data_type. + +Definition init_state := init_global_abstract_data. + +Definition wei := Z. + +Definition addr := int256. + +Definition blocknumber := int256. + +Existing Instance GlobalLayerSpec. + +Section step. + + Context (contract_address : addr). + + Section mstep. + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Definition make_machine_env (caller: addr) + : machine_env state + := {| me_address := contract_address; + me_origin := caller; + me_caller := caller; (* need update after every control-flow transfer *) + me_callvalue := Int256.repr (0); + me_coinbase := coinbase; + me_timestamp := timestamp; + me_number := number; + me_balance := balance; + me_blockhash := blockhash; + (* not implemented *) + me_transfer _ _ _ _ _ := False; + me_callmethod _ _ _ _ _ _ _ _ _ _ := False; + me_log _ _ _ := prev_contract_state; + |}. + + Import MonadNotation. + + Definition lif {A:Type} + (caller : addr) + (cmd : machine_env state -> osT global_abstract_data_type A) + : osT state A := + st <- get;; + let me := make_machine_env caller in + match runStateT (cmd me) st with + | None => mzero + | Some (v, st') => put st' ;; ret v + end. + + (* osT state int256 = + state transformer with option monad, state is state, and value is int256 *) + Print osT. (* = fun D : Type => stateT D option *) + Print stateT. (* (S : Type) (m : Type -> Type) (t : Type) : Type := mkStateT + { runStateT : S -> m (t * S)%type } *) + + Print runStateT. (* takes monad transformer, state, returns m (t * S) *) + + (* How the state of the system changes in a single method call made by player p. *) + Inductive mstep u (st st' : state) : Prop := + | FixedSupplyToken_constructor_step : forall r , runStateT (FixedSupplyToken_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_totalSupply_step : forall r , runStateT (FixedSupplyToken_totalSupply_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_balanceOf_step : forall r tokenOwner, runStateT (FixedSupplyToken_balanceOf_opt tokenOwner (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_transfer_step : forall r toA tokens, runStateT (FixedSupplyToken_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_approve_step : forall r spender tokens, runStateT (FixedSupplyToken_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_transferFrom_step : forall r fromA toA tokens, runStateT (FixedSupplyToken_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_constructor_step : forall r , runStateT (FixedSupplyToken1_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_totalSupply_step : forall r , runStateT (FixedSupplyToken1_totalSupply_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_balanceOf_step : forall r tokenOwner, runStateT (FixedSupplyToken1_balanceOf_opt tokenOwner (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_transfer_step : forall r toA tokens, runStateT (FixedSupplyToken1_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_approve_step : forall r spender tokens, runStateT (FixedSupplyToken1_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken1_transferFrom_step : forall r fromA toA tokens, runStateT (FixedSupplyToken1_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_constructor_step : forall r , runStateT (LiquidityToken_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_mint_step : forall r toA value, runStateT (LiquidityToken_mint_opt toA value (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_burn_step : forall r fromA value, runStateT (LiquidityToken_burn_opt fromA value (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_totalSupply_step : forall r , runStateT (LiquidityToken_totalSupply_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_balanceOf_step : forall r tokenOwner, runStateT (LiquidityToken_balanceOf_opt tokenOwner (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_transfer_step : forall r toA tokens, runStateT (LiquidityToken_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_approve_step : forall r spender tokens, runStateT (LiquidityToken_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | LiquidityToken_transferFrom_step : forall r fromA toA tokens, runStateT (LiquidityToken_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMakerLib_constructor_step : forall r , runStateT (AutomatedMarketMakerLib_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMakerLib_getAmountIn_step : forall r balance amountOut reserve, runStateT (AutomatedMarketMakerLib_getAmountIn_opt balance amountOut reserve (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMakerLib_getBalanceAdjusted_step : forall r balance amountIn, runStateT (AutomatedMarketMakerLib_getBalanceAdjusted_opt balance amountIn (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMakerLib_min_step : forall r x y, runStateT (AutomatedMarketMakerLib_min_opt x y (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_constructor_step : forall r , runStateT (AutomatedMarketMaker_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_mint_step : forall r toA, runStateT (AutomatedMarketMaker_mint_opt toA (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_burn_step : forall r toA, runStateT (AutomatedMarketMaker_burn_opt toA (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_simpleSwap0_step : forall r toA, runStateT (AutomatedMarketMaker_simpleSwap0_opt toA (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_swap_step : forall r amount0Out amount1Out toA, runStateT (AutomatedMarketMaker_swap_opt amount0Out amount1Out toA (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_skim_step : forall r toA, runStateT (AutomatedMarketMaker_skim_opt toA (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_sync_step : forall r , runStateT (AutomatedMarketMaker_sync_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_k_step : forall r , runStateT (AutomatedMarketMaker_k_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_quote0_step : forall r amount0, runStateT (AutomatedMarketMaker_quote0_opt amount0 (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_getAmountOut0_step : forall r amount0In, runStateT (AutomatedMarketMaker_getAmountOut0_opt amount0In (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | AutomatedMarketMaker_getAmountIn0_step : forall r amount0Out, runStateT (AutomatedMarketMaker_getAmountIn0_opt amount0Out (make_machine_env u)) st = Some (r, st') -> mstep u st st' +. + + (* We can compute a new block by letting the players call the contract + in some arbitrary order. *) + Inductive multi_mstep : state -> state -> Prop := + | multi_mstep_reflexive : forall (st : state), multi_mstep st st + | multi_mstep_transitive : forall (st st' st'' : state) u, + multi_mstep st st' -> mstep u st' st'' -> multi_mstep st st''. + + + (* A block is sufficiently synchronous if every player got a chance to submit a + transaction to to. *) + (* TODO, I think this definition is fine for now, but it seems little too clever, + should probably re-state this using some straightforward tracking of the states + we pass through. *) + Definition multi_mstep_synchronous st1 st2 := + forall u, exists st st', + multi_mstep st1 st /\ mstep u st st' /\ multi_mstep st' st2. + + (* Here are a bunch of induction principles inspired by linear temporal logic. *) + + (* Prove that some property P holds "globally", i.e. for each state along a + path. + You can also specify a property Pprev which is known to hold for the + prev_contract_state. If not needed, it can just be True. + *) + Lemma multi_mstep_Global_ind : forall (Pprev P : state -> Prop), + (forall u st st', Pprev prev_contract_state -> P st -> mstep u st st' -> P st') -> + Pprev prev_contract_state -> forall st st', P st -> multi_mstep st st' -> P st'. + Proof. + induction 4; eauto. + Qed. + + (* + (* Prove that P holds "until" Q along a path. + "Until" is a liveness assertion, so we need the synchronicity assumption. *) + Lemma multi_mstep_Until_ind : forall (Pprev P Q : state -> Prop), + (forall p st st', Pprev prev_contract_state -> P st -> In p players + -> mstep p st st' -> (P st' \/ Q st')) -> + Pprev prev_contract_state -> + forall st, + P st -> exists st', multi_mstep st st' -> (P st' \/ Q st'). + Proof. + induction 4; eauto *) + + End mstep. + + + Definition Int256_incr x := Int256.add x Int256.one. + + Inductive bstep (n : blocknumber) : state -> state -> Prop := + | bstep_step : forall coinbase timestamp balance blockhash st st', + multi_mstep coinbase timestamp n balance blockhash st st st' -> + bstep n st st'. + + Inductive multi_bstep : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_reflexive : forall n (st : state), multi_bstep n st n st + | multi_bstep_transitive : forall n n' (st st' st'' : state), + multi_bstep n st n' st' -> bstep (Int256_incr n') st' st'' -> multi_bstep n st (Int256_incr n') st''. + + (* multi_bstep is the step relation without any synchronicity assumption. + This is sufficient to prove some safety properties, but for most interesting + theorems we instead need to use this synchronous version: *) + + Inductive bstep_synch (n : blocknumber) : state -> state -> Prop := + | bstep_synch_step : forall coinbase timestamp balance blockhash st st', + multi_mstep_synchronous coinbase timestamp n balance blockhash st st st' -> + bstep_synch n st st'. + + Inductive multi_bstep_synch : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_synch_reflexive : forall n (st : state), multi_bstep_synch n st n st + | multi_bstep_synch_transitive : forall n n' (st st' st'' : state), + multi_bstep_synch n st n' st' -> bstep_synch (Int256_incr n') st' st'' -> multi_bstep_synch n st (Int256_incr n') st''. + + Lemma multi_bstep_Global_ind : forall (P : state -> Prop), + (forall coinbase timestamp number balance blockhash prev_block u st st', + P prev_block + -> P st + -> mstep coinbase timestamp number balance blockhash prev_block u st st' + -> P st') + -> forall n st n' st', + P st -> multi_bstep n st n' st' -> P st'. + Proof. + induction 3. + - auto. + - inversion H2; subst. + eapply multi_mstep_Global_ind with (st:=st') (prev_contract_state := st'). + + intros. + refine (H _ _ _ _ _ _ _ _ _ _ _ H6); auto. + + apply IHmulti_bstep; assumption. + + apply IHmulti_bstep; assumption. + + exact H3. + Qed. + End step. + + Section DeepSEAGenericProof. + + Lemma Int256Tree_reduce : forall (i: int256) (v: Z) (t: Int256Tree.t Z), Int256Tree.get_default 0%Z i (Int256Tree.set i v t) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gss . + reflexivity. + Qed. + + Lemma Int256Tree_mreduce : forall (i j : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set i v t)) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gss. + reflexivity. + exact H. + Qed. + + Lemma Int256Tree_mireduce : forall (i j k : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + i <> k -> + j <> k -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set k v t)) = + Int256Tree.get_default 0%Z i t. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gso. + reflexivity. + exact H0. + exact H. + Qed. + + Lemma add_sub_inv : forall (i j : Z32), (i + j - i)%Z = j. + Proof. + intros. + omega. + Qed. + End DeepSEAGenericProof. + + Section Proof. + Context (* (strategies : strategy_tuple) *) + (initial_balances : addr -> Z) + (contract_address : int256). + + Context (init_bt init_rt : int256) + (init_coinbase : int256) + (init_timestamp : int256) + (init_number : int256) + (init_blockhash : int256 -> int256) + (pre_init_state init_state : state). + + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Require Import lib.Monad.RunStateTInv. + Require Import lib.ArithInv. + + Definition make_machine_env_wrapped prev_st user := + make_machine_env contract_address coinbase timestamp number balance blockhash prev_st user. + + Lemma make_machine_env_caller_eq : forall st caller, me_caller (make_machine_env_wrapped st caller) = caller. + Proof. auto. Qed. + + Lemma make_machine_env_address_eq : forall st caller, me_address (make_machine_env_wrapped st caller) = contract_address. + Proof. auto. Qed. + +Transparent FixedSupplyToken_constructor_opt. +Transparent FixedSupplyToken_totalSupply_opt. +Transparent FixedSupplyToken_balanceOf_opt. +Transparent FixedSupplyToken_transfer_opt. +Transparent FixedSupplyToken_approve_opt. +Transparent FixedSupplyToken_transferFrom_opt. +Transparent FixedSupplyToken1_constructor_opt. +Transparent FixedSupplyToken1_totalSupply_opt. +Transparent FixedSupplyToken1_balanceOf_opt. +Transparent FixedSupplyToken1_transfer_opt. +Transparent FixedSupplyToken1_approve_opt. +Transparent FixedSupplyToken1_transferFrom_opt. +Transparent LiquidityToken_constructor_opt. +Transparent LiquidityToken_mint_opt. +Transparent LiquidityToken_burn_opt. +Transparent LiquidityToken_totalSupply_opt. +Transparent LiquidityToken_balanceOf_opt. +Transparent LiquidityToken_transfer_opt. +Transparent LiquidityToken_approve_opt. +Transparent LiquidityToken_transferFrom_opt. +Transparent AutomatedMarketMakerLib_constructor_opt. +Transparent AutomatedMarketMakerLib_getAmountIn_opt. +Transparent AutomatedMarketMakerLib_getBalanceAdjusted_opt. +Transparent AutomatedMarketMakerLib_min_opt. +Transparent AutomatedMarketMaker_constructor_opt. +Transparent AutomatedMarketMaker_mint_opt. +Transparent AutomatedMarketMaker_burn_opt. +Transparent AutomatedMarketMaker_simpleSwap0_opt. +Transparent AutomatedMarketMaker_swap_opt. +Transparent AutomatedMarketMaker_skim_opt. +Transparent AutomatedMarketMaker_sync_opt. +Transparent AutomatedMarketMaker_k_opt. +Transparent AutomatedMarketMaker_quote0_opt. +Transparent AutomatedMarketMaker_getAmountOut0_opt. +Transparent AutomatedMarketMaker_getAmountIn0_opt. +Definition FixedSupplyToken_address := (Int256.repr 65587). +Definition FixedSupplyToken1_address := (Int256.repr 65586). +Definition LiquidityToken_address := (Int256.repr 65585). +Definition AutomatedMarketMaker_address := (Int256.repr 65584). + +Definition not_contract_address (addr: int256) := + addr <> FixedSupplyToken_address /\ + addr <> FixedSupplyToken1_address /\ + addr <> LiquidityToken_address /\ + addr <> AutomatedMarketMaker_address. + +Axiom taddr : contract_address = AutomatedMarketMaker_address. + +Ltac rds := +unfold FixedSupplyToken_constructor_opt in *; +unfold FixedSupplyToken_totalSupply_opt in *; +unfold FixedSupplyToken_balanceOf_opt in *; +unfold FixedSupplyToken_transfer_opt in *; +unfold FixedSupplyToken_approve_opt in *; +unfold FixedSupplyToken_transferFrom_opt in *; +unfold FixedSupplyToken1_constructor_opt in *; +unfold FixedSupplyToken1_totalSupply_opt in *; +unfold FixedSupplyToken1_balanceOf_opt in *; +unfold FixedSupplyToken1_transfer_opt in *; +unfold FixedSupplyToken1_approve_opt in *; +unfold FixedSupplyToken1_transferFrom_opt in *; +unfold LiquidityToken_constructor_opt in *; +unfold LiquidityToken_mint_opt in *; +unfold LiquidityToken_burn_opt in *; +unfold LiquidityToken_totalSupply_opt in *; +unfold LiquidityToken_balanceOf_opt in *; +unfold LiquidityToken_transfer_opt in *; +unfold LiquidityToken_approve_opt in *; +unfold LiquidityToken_transferFrom_opt in *; +unfold AutomatedMarketMakerLib_constructor_opt in *; +unfold AutomatedMarketMakerLib_getAmountIn_opt in *; +unfold AutomatedMarketMakerLib_getBalanceAdjusted_opt in *; +unfold AutomatedMarketMakerLib_min_opt in *; +unfold AutomatedMarketMaker_constructor_opt in *; +unfold AutomatedMarketMaker_mint_opt in *; +unfold AutomatedMarketMaker_burn_opt in *; +unfold AutomatedMarketMaker_simpleSwap0_opt in *; +unfold AutomatedMarketMaker_swap_opt in *; +unfold AutomatedMarketMaker_skim_opt in *; +unfold AutomatedMarketMaker_sync_opt in *; +unfold AutomatedMarketMaker_k_opt in *; +unfold AutomatedMarketMaker_quote0_opt in *; +unfold AutomatedMarketMaker_getAmountOut0_opt in *; +unfold AutomatedMarketMaker_getAmountIn0_opt in *; +inv_runStateT; +subst; +inv_arith; +simpl; +try rewrite make_machine_env_caller_eq in *; +try rewrite make_machine_env_address_eq in *; +try rewrite Int256Tree_reduce in *; +try rewrite Int256Tree_mreduce in *; +try rewrite Int256Tree_mireduce in *; +auto. + + Definition compute_k (s: state) := + Z.mul (AutomatedMarketMaker__reserve0 s) (AutomatedMarketMaker__reserve1 s). + + Definition get_balance0 (s: state) (a: addr) := + (Int256Tree.get_default (0%Z) + a (FixedSupplyToken_balances s)). + + Definition get_balance1 (s: state) (a: addr) := + (Int256Tree.get_default 0%Z + a (FixedSupplyToken1_balances s)). + + (* need to transfer balance first, assume token0 transfer first *) + (* should have been approve and then transferFrom, here we are hacking a bit *) + (* prove that after every possible function call, whichever layers, the + k = _reserve0 * _reserve1 is strictly increasing for simpleSwap0: + + newState = simpleSwap0 oldState /\ compute_k newState > compute_k oldState *) + Theorem increasing_k_simpleSwap0 : (forall r r' + (s s' s'' : state) + (trader : addr) + (swapAmount0 : Z32), + (swapAmount0 > 0)%Z -> (* trasferring an amount > 0 *) + ((AutomatedMarketMaker__reserve0 s) > 0)%Z -> + ((AutomatedMarketMaker__reserve1 s) > 0)%Z -> + not_contract_address trader -> + get_balance0 s AutomatedMarketMaker_address = (AutomatedMarketMaker__reserve0 s) -> (* assumes that the reserve is the same as balance at s *) + get_balance1 s AutomatedMarketMaker_address = (AutomatedMarketMaker__reserve1 s) -> + runStateT (FixedSupplyToken_transfer_opt AutomatedMarketMaker_address swapAmount0 (make_machine_env_wrapped s trader)) s = Some (r, s') -> + runStateT (AutomatedMarketMaker_simpleSwap0_opt trader (make_machine_env_wrapped s trader)) s' = Some (r', s'') -> + Z.lt (compute_k s) (compute_k s'')). + Proof. + intros. + unfold get_balance0 in *. + unfold get_balance1 in *. + unfold compute_k. + unfold AutomatedMarketMaker_simpleSwap0_opt in *. + unfold FixedSupplyToken1_transfer_opt in *. + unfold FixedSupplyToken_balanceOf_opt in *. + unfold FixedSupplyToken1_balanceOf_opt in *. + unfold FixedSupplyToken_transfer_opt in *. + inv_runStateT. + rds. + + repeat match goal with + | [ |- context[AutomatedMarketMaker__reserve1 ?X]] => remember (AutomatedMarketMaker__reserve1 X) as R1 + end. + repeat match goal with + | [ |- context[AutomatedMarketMaker__reserve0 ?X]] => remember (AutomatedMarketMaker__reserve0 X) as R0 + end. + match goal with + | [ |- context[Z.mul (Z.add ?X _) _]] => replace X with R0 (* rewrite -> H1 *) + end. + match goal with + | [ |- context[Z.sub ?X _]] => replace X with R1 + end. + remember swapAmount0 as delta0. + +End Proof. diff --git a/minic_unittests/contracts/amm/amm/prf_int.v b/minic_unittests/contracts/amm/amm/prf_int.v new file mode 100644 index 0000000..484022b --- /dev/null +++ b/minic_unittests/contracts/amm/amm/prf_int.v @@ -0,0 +1,631 @@ +Require Import amm.DataTypes. +Require Import amm.DataTypeOps. +Require Import amm.LayerAMM. +Require Import lib.Monad.StateMonadOption. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import ZArith. +Require Import core.HyperTypeInst. +(* for `hashvalue`, maybe should move that to some different file. *) +Require Import backend.MachineModel. + +(* PROOFS: + 2. balance tracks reserve + 3. increasing k -> algebra + 4. no path independence, splitting trades is strictly more expensive -> algebra, appendix D + 5. liquidity token economics is right, i.e., UNI-V2 tracks the actual share of pool --> actually, where is the trading of UNI-V2 Pair tokens? What is its value? + 6. MOD: slippage control + 7. ... cost of manipulation (hard) +*) + +(* global_abstract_data_type *) + +Definition state := global_abstract_data_type. + +Definition init_state := init_global_abstract_data. + +Definition wei := Z. + +Definition addr := int256. + +Definition blocknumber := int256. + +Existing Instance GlobalLayerSpec. + +Section step. + + Context (token0_address token1_address amm_address : addr). + + Section mstep. + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Definition make_machine_env (caller: addr) + : machine_env state + := {| me_address := amm_address; + me_origin := caller; + me_caller := caller; + me_callvalue := Int256.repr (0); + me_coinbase := coinbase; + me_timestamp := timestamp; + me_number := number; + me_balance := balance; + me_blockhash := blockhash; + (* not implemented *) + me_transfer _ _ _ _ _ := False; + me_callmethod _ _ _ _ _ _ _ _ _ _ := False; + me_log _ _ _ := prev_contract_state; + |}. + + Import MonadNotation. + + Definition lif {A:Type} + (caller : addr) + (cmd : machine_env state -> osT global_abstract_data_type A) + : osT state A := + st <- get;; + let me := make_machine_env caller in + match runStateT (cmd me) st with + | None => mzero + | Some (v, st') => put st' ;; ret v + end. + + Definition lifted_simpleSwap0 := + fun (caller: addr) (callvalue: wei) (toA: addr) => lif caller + (AutomatedMarketMaker_simpleSwap0_opt toA). + + Definition lifted_LiquidityToken_mint := + fun (caller: addr) (callvalue: wei) (toA: addr) (value: Z) => lif caller + (LiquidityToken_mint_opt toA value). + + (* osT state int256 = + state transformer with option monad, state is state, and value is int256 *) + Print osT. (* = fun D : Type => stateT D option *) + Print stateT. (* (S : Type) (m : Type -> Type) (t : Type) : Type := mkStateT + { runStateT : S -> m (t * S)%type } *) + + Print runStateT. (* takes monad transformer, state, returns m (t * S) *) + + (* How the state of the system changes in a single method call made by player p. *) + (* for simplicity we only model functions that modify states and can be called by user *) + Inductive mstep u (st st' : state) : Prop := + (* fixed supply tokens *) + | FixedSupplyToken0_transfer_step : + forall toA tokens r, runStateT (FixedSupplyToken_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | FixedSupplyToken1_transfer_step : + forall toA tokens r, runStateT (FixedSupplyToken1_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | FixedSupplyToken0_approve_step : + forall spender tokens r, runStateT (FixedSupplyToken_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | FixedSupplyToken1_approve_step : + forall spender tokens r, runStateT (FixedSupplyToken1_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | FixedSupplyToken0_transferFrom_step : + forall fromA toA tokens r, runStateT (FixedSupplyToken_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | FixedSupplyToken1_transferFrom_step : + forall fromA toA tokens r, runStateT (FixedSupplyToken1_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + (* amm *) + | AutomatedMarketMaker_mint_step : + forall toA r, runStateT (AutomatedMarketMaker_mint_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | AutomatedMarketMaker_burn_step : + forall toA r, runStateT (AutomatedMarketMaker_burn_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + | AutomatedMarketMaker_simpleSwap0_step : + forall toA r, runStateT (AutomatedMarketMaker_simpleSwap0_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep u st st' + . + + (* How the state of the system changes in a single method call made by player p. *) + (* for simplicity we only model functions that modify states and can be called by user *) + Inductive mstep_amm u (st st' : state) : Prop := + (* amm *) + | AutomatedMarketMaker_mint_ammstep : + forall toA r, runStateT (AutomatedMarketMaker_mint_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep_amm u st st' + | AutomatedMarketMaker_burn_ammstep : + forall toA r, runStateT (AutomatedMarketMaker_burn_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep_amm u st st' + | AutomatedMarketMaker_simpleSwap0_ammstep : + forall toA r, runStateT (AutomatedMarketMaker_simpleSwap0_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep_amm u st st' + | AutomatedMarketMaker_skim_ammstep : + forall toA r, runStateT (AutomatedMarketMaker_skim_opt toA (make_machine_env u)) st = Some (r, st') + -> mstep_amm u st st' + | AutomatedMarketMaker_sync_ammstep : + forall r, runStateT (AutomatedMarketMaker_sync_opt (make_machine_env u)) st = Some (r, st') + -> mstep_amm u st st' + . + + (* We can compute a new block by letting the players call the contract + in some arbitrary order. *) + Inductive multi_mstep : state -> state -> Prop := + | multi_mstep_reflexive : forall (st : state), multi_mstep st st + | multi_mstep_transitive : forall (st st' st'' : state) u, + multi_mstep st st' -> mstep u st' st'' -> multi_mstep st st''. + + + (* A block is sufficiently synchronous if every player got a chance to submit a + transaction to to. *) + (* TODO, I think this definition is fine for now, but it seems little too clever, + should probably re-state this using some straightforward tracking of the states + we pass through. *) + Definition multi_mstep_synchronous st1 st2 := + forall u, exists st st', + multi_mstep st1 st /\ mstep u st st' /\ multi_mstep st' st2. + + (* Here are a bunch of induction principles inspired by linear temporal logic. *) + + (* Prove that some property P holds "globally", i.e. for each state along a + path. + You can also specify a property Pprev which is known to hold for the + prev_contract_state. If not needed, it can just be True. + *) + Lemma multi_mstep_Global_ind : forall (Pprev P : state -> Prop), + (forall u st st', Pprev prev_contract_state -> P st -> mstep u st st' -> P st') -> + Pprev prev_contract_state -> forall st st', P st -> multi_mstep st st' -> P st'. + Proof. + induction 4; eauto. + Qed. + + (* + (* Prove that P holds "until" Q along a path. + "Until" is a liveness assertion, so we need the synchronicity assumption. *) + Lemma multi_mstep_Until_ind : forall (Pprev P Q : state -> Prop), + (forall p st st', Pprev prev_contract_state -> P st -> In p players + -> mstep p st st' -> (P st' \/ Q st')) -> + Pprev prev_contract_state -> + forall st, + P st -> exists st', multi_mstep st st' -> (P st' \/ Q st'). + Proof. + induction 4; eauto *) + + End mstep. + + + Definition Int256_incr x := Int256.add x Int256.one. + + Inductive bstep (n : blocknumber) : state -> state -> Prop := + | bstep_step : forall coinbase timestamp balance blockhash st st', + multi_mstep coinbase timestamp n balance blockhash st st st' -> + bstep n st st'. + + Inductive multi_bstep : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_reflexive : forall n (st : state), multi_bstep n st n st + | multi_bstep_transitive : forall n n' (st st' st'' : state), + multi_bstep n st n' st' -> bstep (Int256_incr n') st' st'' -> multi_bstep n st (Int256_incr n') st''. + + (* multi_bstep is the step relation without any synchronicity assumption. + This is sufficient to prove some safety properties, but for most interesting + theorems we instead need to use this synchronous version: *) + + Inductive bstep_synch (n : blocknumber) : state -> state -> Prop := + | bstep_synch_step : forall coinbase timestamp balance blockhash st st', + multi_mstep_synchronous coinbase timestamp n balance blockhash st st st' -> + bstep_synch n st st'. + + Inductive multi_bstep_synch : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_synch_reflexive : forall n (st : state), multi_bstep_synch n st n st + | multi_bstep_synch_transitive : forall n n' (st st' st'' : state), + multi_bstep_synch n st n' st' -> bstep_synch (Int256_incr n') st' st'' -> multi_bstep_synch n st (Int256_incr n') st''. + + Lemma multi_bstep_Global_ind : forall (P : state -> Prop), + (forall coinbase timestamp number balance blockhash prev_block u st st', + P prev_block + -> P st + -> mstep coinbase timestamp number balance blockhash prev_block u st st' + -> P st') + -> forall n st n' st', + P st -> multi_bstep n st n' st' -> P st'. + Proof. + induction 3. + - auto. + - inversion H2; subst. + eapply multi_mstep_Global_ind with (st:=st') (prev_contract_state := st'). + + intros. + refine (H _ _ _ _ _ _ _ _ _ _ _ H6); auto. + + apply IHmulti_bstep; assumption. + + apply IHmulti_bstep; assumption. + + exact H3. + Qed. +End step. + +Section DeepSEAGenericProof. + + Lemma Int256Tree_reduce : forall (i: int256) (v: Z) (t: Int256Tree.t Z), Int256Tree.get_default 0%Z i (Int256Tree.set i v t) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gss . + reflexivity. + Qed. + + Lemma Int256Tree_mreduce : forall (i j : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set i v t)) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gss. + reflexivity. + exact H. + Qed. + + Lemma Int256Tree_mireduce : forall (i j k : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + i <> k -> + j <> k -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set k v t)) = + Int256Tree.get_default 0%Z i t. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gso. + reflexivity. + exact H0. + exact H. + Qed. + + Lemma add_sub_inv : forall (i j : Z32), (i + j - i)%Z = j. + Proof. + intros. + omega. + Qed. + + (* Lemmas on Inequalities over Z *) + Lemma ZInEq_ladd : (forall (i j k : Z32), (i > j + k <-> i - k > j)%Z). + Proof. + intros. omega. + Qed. + Hint Rewrite ZInEq_ladd : ZInEq. + + Axiom ZInEq_multpos : (forall (a b c : Z32), c > 0 -> a > b <-> a * c > b * c)%Z. + Hint Rewrite ZInEq_multpos : ZInEq. + + Axiom ZInEq_ldiv : (forall (a b c : Z32), b > 0 -> a / b > c <-> a > b * c)%Z. + + Axiom ZInEq_denadd : (forall (a b c : Z32), c > 0 -> a / b > a / (b + c))%Z. + + Lemma increasing_k_math : forall (i j k : Z), (i > 0 -> j > 0 -> k > 0 -> + i * j < (i + k) * (j - (k * 997 * j) / (i * 1000 + k * 997)))%Z. + Proof. + From Coq Require Import ZArith Psatz. + Open Scope Z_scope. + simpl. Fail lia. Print Lia.lia. Print Lia.zchecker_no_abstract. + intros. + (* + psatz Z 100. + lia. + + assert (forall (v v' v'' : Z), v'' > 0 -> v < v' -> v * v'' < v' * v''). + intros. + pose (H19 := H16 (v * v'') (v' * v'')); clearbody H19. + apply H19. + unfold Z.sub. + apply mul_sub_distr. + pose (H17 := H16 (i * j) ((i + k) * (j - k * 997 * j / (i * 1000 + k * 997))) (i * 1000 + k * 997)); clearbody H17. + *) + assert (0 < (i * 1000 + k * 997)). omega. + (* multiplied both sides by the denominator of the divisor and needed to use comparisons between "(W / V) * V" and W *) + pose (Zmult_nc:=(Zmult_lt_compat_r (i * j) ((i + k) * (j - k * 997 * j / (i * 1000 + k * 997))) (i * 1000 + k * 997) H2)); clearbody Zmult_nc. + Admitted. +End DeepSEAGenericProof. + +Section AMMProof. + Context (token0_address token1_address amm_address : addr). + + Context (* (strategies : strategy_tuple) *) + (initial_balances : addr -> Z) + (contract_address : int256). + + Context (init_bt init_rt : int256) + (init_coinbase : int256) + (init_timestamp : int256) + (init_number : int256) + (init_blockhash : int256 -> int256) + (pre_init_state init_state : state). + + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Require Import lib.Monad.RunStateTInv. + Require Import lib.ArithInv. + + Definition compute_k (s: state) := + Z.mul (AutomatedMarketMaker__reserve0 s) (AutomatedMarketMaker__reserve1 s). + + Definition get_balance0 (s: state) (a: addr) := + (Int256Tree.get_default (0%Z) + a (FixedSupplyToken_balances s)). + + Definition get_balance1 (s: state) (a: addr) := + (Int256Tree.get_default 0%Z + a (FixedSupplyToken1_balances s)). + + Definition make_machine_env_wrapped prev_st user := + make_machine_env amm_address coinbase timestamp number balance blockhash prev_st user. + + Lemma make_machine_env_caller_eq : forall st caller, me_caller (make_machine_env_wrapped st caller) = caller. + Proof. auto. Qed. + + Lemma make_machine_env_address_eq : forall st caller, me_address (make_machine_env_wrapped st caller) = amm_address. + Proof. auto. Qed. + +Transparent FixedSupplyToken_constructor_opt. +Transparent FixedSupplyToken_totalSupply_opt. +Transparent FixedSupplyToken_balanceOf_opt. +Transparent FixedSupplyToken_transfer_opt. +Transparent FixedSupplyToken_approve_opt. +Transparent FixedSupplyToken_transferFrom_opt. +Transparent FixedSupplyToken1_constructor_opt. +Transparent FixedSupplyToken1_totalSupply_opt. +Transparent FixedSupplyToken1_balanceOf_opt. +Transparent FixedSupplyToken1_transfer_opt. +Transparent FixedSupplyToken1_approve_opt. +Transparent FixedSupplyToken1_transferFrom_opt. +Transparent LiquidityToken_constructor_opt. +Transparent LiquidityToken_mint_opt. +Transparent LiquidityToken_burn_opt. +Transparent LiquidityToken_totalSupply_opt. +Transparent LiquidityToken_balanceOf_opt. +Transparent LiquidityToken_transfer_opt. +Transparent LiquidityToken_approve_opt. +Transparent LiquidityToken_transferFrom_opt. +Transparent AutomatedMarketMakerLib_constructor_opt. +Transparent AutomatedMarketMakerLib_getAmountIn_opt. +Transparent AutomatedMarketMakerLib_getBalanceAdjusted_opt. +Transparent AutomatedMarketMakerLib_min_opt. +Transparent AutomatedMarketMaker_constructor_opt. +Transparent AutomatedMarketMaker_mint_opt. +Transparent AutomatedMarketMaker_burn_opt. +Transparent AutomatedMarketMaker_simpleSwap0_opt. +Transparent AutomatedMarketMaker_swap_opt. +Transparent AutomatedMarketMaker_skim_opt. +Transparent AutomatedMarketMaker_sync_opt. +Transparent AutomatedMarketMaker_k_opt. +Transparent AutomatedMarketMaker_quote0_opt. +Transparent AutomatedMarketMaker_getAmountOut0_opt. +Transparent AutomatedMarketMaker_getAmountIn0_opt. +Definition FixedSupplyToken_address := (Int256.repr 65587). +Definition FixedSupplyToken1_address := (Int256.repr 65586). +Definition LiquidityToken_address := (Int256.repr 65585). +Definition AutomatedMarketMaker_address := (Int256.repr 65584). + +Ltac rds := +unfold FixedSupplyToken_constructor_opt in *; +unfold FixedSupplyToken_totalSupply_opt in *; +unfold FixedSupplyToken_balanceOf_opt in *; +unfold FixedSupplyToken_transfer_opt in *; +unfold FixedSupplyToken_approve_opt in *; +unfold FixedSupplyToken_transferFrom_opt in *; +unfold FixedSupplyToken1_constructor_opt in *; +unfold FixedSupplyToken1_totalSupply_opt in *; +unfold FixedSupplyToken1_balanceOf_opt in *; +unfold FixedSupplyToken1_transfer_opt in *; +unfold FixedSupplyToken1_approve_opt in *; +unfold FixedSupplyToken1_transferFrom_opt in *; +unfold LiquidityToken_constructor_opt in *; +unfold LiquidityToken_mint_opt in *; +unfold LiquidityToken_burn_opt in *; +unfold LiquidityToken_totalSupply_opt in *; +unfold LiquidityToken_balanceOf_opt in *; +unfold LiquidityToken_transfer_opt in *; +unfold LiquidityToken_approve_opt in *; +unfold LiquidityToken_transferFrom_opt in *; +unfold AutomatedMarketMakerLib_constructor_opt in *; +unfold AutomatedMarketMakerLib_getAmountIn_opt in *; +unfold AutomatedMarketMakerLib_getBalanceAdjusted_opt in *; +unfold AutomatedMarketMakerLib_min_opt in *; +unfold AutomatedMarketMaker_constructor_opt in *; +unfold AutomatedMarketMaker_mint_opt in *; +unfold AutomatedMarketMaker_burn_opt in *; +unfold AutomatedMarketMaker_simpleSwap0_opt in *; +unfold AutomatedMarketMaker_swap_opt in *; +unfold AutomatedMarketMaker_skim_opt in *; +unfold AutomatedMarketMaker_sync_opt in *; +unfold AutomatedMarketMaker_k_opt in *; +unfold AutomatedMarketMaker_quote0_opt in *; +unfold AutomatedMarketMaker_getAmountOut0_opt in *; +unfold AutomatedMarketMaker_getAmountIn0_opt in *; +inv_runStateT; +subst; +inv_arith; +simpl in *; +try rewrite make_machine_env_caller_eq in *; +try rewrite make_machine_env_address_eq in *; +try rewrite Int256Tree_reduce in *; +try rewrite Int256Tree_mreduce in *; +try rewrite Int256Tree_mireduce in *; +auto. + + (* need to transfer balance first, assume token0 transfer first *) + (* should have been approve and then transferFrom, here we are hacking a bit *) + (* prove that after every possible function call, whichever layers, the + k = _reserve0 * _reserve1 is strictly increasing for simpleSwap0: + + newState = simpleSwap0 oldState /\ compute_k newState > compute_k oldState *) + Theorem increasing_k_simpleSwap0 : (forall r r' + (s s' s'' : state) + (trader : addr) + (swapAmount0 : Z32), + (swapAmount0 > 0)%Z -> (* trasferring an amount > 0 *) + ((AutomatedMarketMaker__reserve0 s) > 0)%Z -> + ((AutomatedMarketMaker__reserve1 s) > 0)%Z -> + amm_address <> trader -> + get_balance0 s amm_address = (AutomatedMarketMaker__reserve0 s) -> (* assumes that the reserve is the same as balance at s *) + get_balance1 s amm_address = (AutomatedMarketMaker__reserve1 s) -> + runStateT (FixedSupplyToken_transfer_opt amm_address swapAmount0 (make_machine_env_wrapped s trader)) s = Some (r, s') -> + runStateT (AutomatedMarketMaker_simpleSwap0_opt trader (make_machine_env_wrapped s amm_address)) s' = Some (r', s'') -> + Z.lt (compute_k s) (compute_k s'')). + Proof. + intros. + unfold get_balance0 in H0. + unfold get_balance1 in H1. + unfold AutomatedMarketMaker_simpleSwap0_opt in *. + unfold FixedSupplyToken1_transfer_opt in *. + unfold FixedSupplyToken_balanceOf_opt in *. + unfold FixedSupplyToken1_balanceOf_opt in *. + unfold FixedSupplyToken_transfer_opt in *. + inv_runStateT. + rewrite -> make_machine_env_caller_eq in *. + rewrite -> make_machine_env_address_eq in *. + subst. + inv_arith. + unfold compute_k. + simpl. + + rewrite Int256Tree_reduce in *. + rewrite Int256Tree_mreduce in *. + repeat match goal with + | [ |- context[AutomatedMarketMaker__reserve1 ?X]] => remember (AutomatedMarketMaker__reserve1 X) as R1 + end. + repeat match goal with + | [ |- context[AutomatedMarketMaker__reserve0 ?X]] => remember (AutomatedMarketMaker__reserve0 X) as R0 + end. + match goal with + | [ |- context[Z.mul (Z.add ?X _) _]] => replace X with R0 (* rewrite -> H1 *) + end. + match goal with + | [ |- context[Z.sub ?X _]] => replace X with R1 + end. + remember swapAmount0 as delta0. + + pose (IKM := increasing_k_math R0 R1 delta0 H0 H1 H); clearbody IKM. + rewrite add_sub_inv. + exact IKM. + exact H2. + Qed. + +Lemma balance_tracks_reserve_inv : forall u s s', + u <> amm_address -> + get_balance0 s amm_address = (AutomatedMarketMaker__reserve0 s) -> + get_balance1 s amm_address = (AutomatedMarketMaker__reserve1 s) -> + mstep_amm amm_address coinbase timestamp number balance blockhash prev_contract_state u s s' -> + get_balance0 s amm_address = (AutomatedMarketMaker__reserve0 s') /\ + get_balance1 s amm_address = (AutomatedMarketMaker__reserve1 s'). +Proof. + intros u s s' Hun Hps0 Hps1 Hstp. + unfold get_balance0 in *. + unfold get_balance1 in *. + destruct Hstp; rds. (* FIXME: rds should be able to automatically solve all the goals if we have + machine_env updating implemented *) +Admitted. + +Axiom anything : False. + +(* splitting trades is strictly more expensive, if you split one trade into two trades with combined + same amount, then you swapped for less alternative tokens *) +Theorem path_dependency : forall r r' rd rd' rd'' rd''' + (s s' s'' sd' sd'' sd''' sd'''' : state) + (trader : addr) + (swapAmount0 : Z32) + (swapAmount0D0 swapAmount0D1 : Z32), + (swapAmount0D0 > 0)%Z -> + (swapAmount0D1 > 0)%Z -> + (swapAmount0D0 + swapAmount0D1 = swapAmount0)%Z -> + (swapAmount0 > 0)%Z -> (* trasferring an amount > 0 *) + ((AutomatedMarketMaker__reserve0 s) > 0)%Z -> + ((AutomatedMarketMaker__reserve1 s) > 0)%Z -> + amm_address <> trader -> + get_balance0 s amm_address = (AutomatedMarketMaker__reserve0 s) -> (* assumes that the reserve is the same as balance at s *) + get_balance1 s amm_address = (AutomatedMarketMaker__reserve1 s) -> + (* swapping in one trade *) + runStateT (FixedSupplyToken0_transfer_opt amm_address swapAmount0 (make_machine_env_wrapped s trader)) s = Some (r, s') -> + runStateT (AutomatedMarketMaker_simpleSwap0_opt trader (make_machine_env_wrapped s amm_address)) s' = Some (r', s'') -> + (* swapping in two trades *) + runStateT (FixedSupplyToken0_transfer_opt amm_address swapAmount0D0 (make_machine_env_wrapped s trader)) s = Some (rd, sd') -> + runStateT (AutomatedMarketMaker_simpleSwap0_opt trader (make_machine_env_wrapped s amm_address)) sd' = Some (rd', sd'') -> + runStateT (FixedSupplyToken0_transfer_opt amm_address swapAmount0D1 (make_machine_env_wrapped s trader)) sd'' = Some (rd'', sd''') -> + runStateT (AutomatedMarketMaker_simpleSwap0_opt trader (make_machine_env_wrapped s amm_address)) sd''' = Some (rd''', sd'''') -> + (* the ending balance is different *) + (r' > rd' + rd''')%Z. +Proof. + (* generic tactic *) + intros. + unfold get_balance0 in *. + unfold get_balance1 in *. + Time rds. (* finish in 182 secs, or 3m2s *) + rewrite !H6. rewrite !H7. + remember (AutomatedMarketMaker__reserve0 m121) as a. + remember (AutomatedMarketMaker__reserve1 m121) as b. + remember (swapAmount0D0) as c. + remember (swapAmount0D1) as d. + rewrite !add_sub_inv. (* need inequality rewriting rules *) + + assert (d * 997 * R1 / (R0 * 1000 + d * 997) > d * 997 * R1 / (R0 * 1000 + d * 997 + d' * 997))%Z. + assert (d' * 997 > 0)%Z. omega. + match goal with + | [ |- context[Z.gt (Z.div ?A ?B) (Z.div ?A (Z.add ?B ?C))]] => pose (HZInEq_denadd := ZInEq_denadd A B C H31); clearbody HZInEq_denadd + end. + exact HZInEq_denadd. + apply H31. + rewrite Zgt_trans. + nia. + unfold AutomatedMarketMaker_simpleSwap0_opt in *. + unfold FixedSupplyToken1_transfer_opt in *. + unfold FixedSupplyToken0_balanceOf_opt in *. + unfold FixedSupplyToken1_balanceOf_opt in *. + unfold FixedSupplyToken0_transfer_opt in *. + + Opaque bind ret get gets put modify guard mzero. + + Import MonadNotation. + inv_runStateT1 H13. + subst. + inv_arith. + simpl in *. + + inv_runStateT1 H12. + subst. + autorewrite with updates in *. + rewrite !Int256Tree_reduce in *. + simpl in *. + inv_arith. + + inv_runStateT1 H11. + subst. + simpl in *. + autorewrite with updates in *. + inv_arith. + rewrite !(@Int256Tree_Properties.get_default_so _ (0%Z) a25 a) in * by (exact H22). + rewrite !Int256Tree_Properties.get_default_ss in *. + + inv_runStateT1 H9. + subst. + simpl in *. + autorewrite with updates in *. + inv_arith. + + inv_runStateT1 H8. + subst. + simpl in *. + autorewrite with updates in *. + rewrite !Int256Tree_Properties.get_default_ss in *. + + inv_runStateT1 H10. + subst. + simpl in *. + autorewrite with updates in *. + rewrite !Int256Tree_Properties.get_default_ss in *. + + (* These lines are to test if the resulting proof term can be checked reasonably quickly, but it seems it works. *) + (* destruct anything. + Time Qed. *) + +Admitted. diff --git a/minic_unittests/contracts/auction/auction.ds b/minic_unittests/contracts/auction/auction.ds new file mode 100755 index 0000000..df99c48 --- /dev/null +++ b/minic_unittests/contracts/auction/auction.ds @@ -0,0 +1,55 @@ +(* TODO: + - implement withdrawals + *) + +object signature AuctionInterface = { + initialize : uint -> unit; + const getbid : unit -> uint; + const getbidder : unit -> address; + const getchair : unit -> address; + const getdeadline : unit -> uint; + bid : unit -> bool; +} + +object OpenAuction : AuctionInterface { + (* initial values don't get used *) + let _bid : uint := 0u0 + let _bidder : address := address(0) + let _chair : address := address(0) + let _deadline : uint := 0u0 + let withdrawals : mapping[address] uint := mapping_init + + (* placeholder for constructor *) + let initialize deadline = + _bid := 0u0; + _bidder := msg_sender; + _chair := msg_sender; + _deadline := deadline; + () + + (* getter functions *) + let getbid () = _bid + let getbidder () = _bidder + let getchair () = _chair + let getdeadline () = _deadline + + (* place bid *) + let bid () = + let bidder = _bidder in + let bid = _bid in + let deadline = _deadline in + let withdrawal = withdrawals[bidder] in + + assert ((msg_sender <> bidder ) /\ + (msg_value > bid ) /\ + (block_number < deadline)); + + withdrawals[bidder] := withdrawal + bid; + _bidder := msg_sender; + _bid := msg_value; + true +} + +layer AUCTION = { + auction = OpenAuction +} diff --git a/minic_unittests/contracts/auction/auction/FunctionalCorrectness.v b/minic_unittests/contracts/auction/auction/FunctionalCorrectness.v new file mode 100755 index 0000000..cf0f156 --- /dev/null +++ b/minic_unittests/contracts/auction/auction/FunctionalCorrectness.v @@ -0,0 +1,72 @@ +Require Import backend.MachineModel. +Require Import cclib.Integers. +Require Import LayerAUCTION. +Require Import backend.Values. +Require Import backend.BuiltinSemantics. +Require Import DeepSpec.lib.Monad.StateMonadOption. +Require Import DataTypeOps. + + + +Obligation Tactic := try eexists; simpl; eauto; try congruence. +Program Definition simple_machine_env + (address origin caller callvalue + coinbase timestamp number : int256) : machine_env := + {| me_query q := + match q with + | Qcall0 Baddress => Vint address + | Qcall0 Borigin => Vint origin + | Qcall0 Bcaller => Vint caller + | Qcall0 Bcallvalue => Vint callvalue + | Qcall0 Bcoinbase => Vint coinbase + | Qcall0 Btimestamp => Vint timestamp + | Qcall0 Bnumber => Vint number + | Qcall1 Bsha_1 v => Vunit (* todo. *) + | Qcall1 Bbalance v => Vunit (* todo. *) + | Qcall1 Bblockhash v => Vunit (* todo. *) + | Qcall2 Bsha_2 (Vptr a1p) (Vint a2i) => Vptr (Ihash a1p a2i) + | Qcall2 Bsha_2 _ _ => Vunit (* todo. *) + end ; + me_transfer _ _ _ _ := True; + me_callmethod _ _ _ _ _ _ _ := True + |}. + +Definition from_option {A:Type} (default:A) (x:option A) :A := + match x with + | None => default + | Some y => y + end. + +Require Import ZArith. +Require Import cclib.Maps. + +Section WithMem. + +Import core.MemoryModel. + +Context {HmemOps: MemoryModelOps mem}. + +Section AuctionVerification. + + Parameter owner : int256. + Parameter contract_address : int256. + Parameter deadline : int256. + + Context (init_coinbase + init_timestamp + init_number : int256). + + Transparent OpenAuction_initialize_opt. + Transparent HyperTypeInst.builtin0_caller_impl HyperTypeInst.builtin0_callvalue_impl HyperTypeInst.builtin0_coinbase_impl HyperTypeInst.builtin0_timestamp_impl HyperTypeInst.builtin0_number_impl. + + Definition OpenAuction_initial := + Eval cbv -[Int256Tree.empty Int256.zero] + in + from_option init_global_abstract_data + (execStateT (OpenAuction_initialize_opt + deadline + (simple_machine_env contract_address owner owner Int256.zero + init_coinbase init_timestamp init_number)) + init_global_abstract_data). + + \ No newline at end of file diff --git a/minic_unittests/contracts/olive/Olive.ds b/minic_unittests/contracts/olive/Olive.ds new file mode 100755 index 0000000..f8f42a3 --- /dev/null +++ b/minic_unittests/contracts/olive/Olive.ds @@ -0,0 +1,234 @@ +(* This contract is inspired by olive/contracts/Olive.sol , a very vanilla ERC20 token contract. *) + +object signature OwnableInterface = { + constructor : unit -> unit; + onlyOwner : unit -> unit; + transferOwnership : address -> unit +} + +event + | OwnershipTransferred (previousOwner : address indexed) (newOwner: address indexed) + +object Ownable : OwnableInterface { + let owner : address := address(0) + + let constructor () = + owner := msg_sender + + let onlyOwner () = + let a = owner in + assert(a = msg_sender) + + let transferOwnership(newOwner) = + (* assert (newOwner != 0); *) + let b = owner in + emit OwnershipTransferred(b, newOwner); + owner := newOwner +} + +layer OWNABLE : [ { } ] { ownable : OwnableInterface } = { + ownable = Ownable +} + +object signature PausableInterface = { + whenPaused : unit -> unit; + whenNotPaused : unit -> unit; + pause : unit -> unit; + unpause : unit -> unit +} + +event + | Pause(v : bool) + | Unpause(v : bool) + +object Pausable (ownable : OwnableInterface) : PausableInterface { + let paused : bool := false + + let whenPaused () = + let a = paused in + assert (a = true) + + let whenNotPaused () = + let b = paused in + assert (b = false) + + let pause () = + ownable.onlyOwner(); + paused := true; + let a = paused in + emit Pause(a) + + let unpause () = + ownable.onlyOwner(); + paused := false; + let a = paused in + emit Unpause(a) + +} + +layer PAUSABLE : [{ ownable : OwnableInterface }] + { pausable : PausableInterface } += { + pausable = Pausable +} + +event + | Transfer(from : address indexed) (to_: address indexed) (value: uint) + | Approval(owner: address indexed) (spender: address indexed) (value: uint) + +object signature LockableInterface = { + increaseLockedValue : address * uint -> unit; + const isLocked : uint -> bool; + lockCheck : address * uint * uint -> bool; + lockAddress : address -> unit; + unlockAddress : address -> unit; + unlock : unit -> unit +} + +object Lockable (ownable : OwnableInterface): LockableInterface { + let _balanceOfLocked : mapping[address] uint := mapping_init + let _addressLocked : mapping[address] bool := mapping_init + let _unlocktime : uint := 0u1527868800 + let _manualUnlock : bool := false + + let increaseLockedValue(addr, n) = + let a = _balanceOfLocked[addr] in + _balanceOfLocked[addr] := a + n + + let isLocked now = + let unlock_time = _unlocktime in + let manual_unlock = _manualUnlock in + ((now < unlock_time) /\ (!manual_unlock)) + + let lockCheck (from, val, now) = + (* assert (_addressLocked[from] == false); *) + let unlock_time = _unlocktime in + let manual_unlock = _manualUnlock in + if ((now < unlock_time) /\ (!manual_unlock)) + then begin + (* assert(val <= (balances[from] - balanceOfLocked[from])); *) + true + end else + false + + let lockAddress addr = + _addressLocked[addr] := true + + let unlockAddress addr = + _addressLocked[addr] := false + + let unlock () = + ownable.onlyOwner(); + let manual_unlock = _manualUnlock in + assert(!manual_unlock); + _manualUnlock := true + +} + +layer LOCKABLE : [{ ownable : OwnableInterface }] + { lockable : LockableInterface } += { + lockable = Lockable +} + +(* we cannot have "to" which is a token that could cause parse error; need to mention in reference *) + +object signature ERC20Interface = { + constructor : unit -> unit; + const balanceOf : address -> uint; + transfer : address * uint * uint -> bool; + allowance : address * address -> uint; + transferFrom : address * address * uint * uint -> bool; + approve : address * uint -> bool; + + increaseApproval : address * uint -> bool; + decreaseApproval : address * uint -> bool +} + +object Token (pausable : PausableInterface, ownable : OwnableInterface, lockable : LockableInterface) : ERC20Interface { + let _totalSupply : uint := 0u10000000000 + let _balances : mapping[address] uint := mapping_init + let _allowed : mapping[address] mapping[address] uint := mapping_init + + let constructor () = + let a = _totalSupply in + _balances[msg_sender] := a + + let balanceOf addr = + _balances[addr] + + let transfer (to_, value, now) = + let check = lockable.lockCheck(msg_sender, value, now) in + assert(check <> false); + pausable.whenNotPaused(); + let from_ = msg_sender in + let sender_bal = _balances[from_] in + let receiver_bal = _balances[to_] in + assert(to_ <> address(0)); + assert(value <= sender_bal); + assert (sender_bal >= value); + _balances[from_] := sender_bal - value; + (* todo: add suitable overflow assertion here. *) + _balances[to_] := receiver_bal + value; + emit Transfer(msg_sender, to_, value); + true + + let allowance (owner, spender) = + let res = _allowed[owner][spender] in + res + + let transferFrom (from_, to_, value, now) = + let check = lockable.lockCheck(from_, value, now) in + assert(check <> false); + pausable.whenNotPaused(); + let from_bal = _balances[from_] in + let from_allow = _allowed[from_][msg_sender] in + let to_bal = _balances[to_] in + assert(to_ <> address(0)); + assert(value <= from_bal); + assert(value <= from_allow); + _balances[from_] := from_bal - value; + (* todo: overflow check for the addition. *) + _balances[to_] := to_bal + value; + _allowed[from_][msg_sender] := from_allow- (value); + emit Transfer(from_, to_, value); + true + + let approve (spender, value) = + (* pausable.whenNotPaused(); *) + _allowed[msg_sender][spender] := value; + emit Approval(msg_sender, spender, value); + true + + + let increaseApproval(spender, addedValue) = + pausable.whenNotPaused(); + (* todo: overflow check *) + let allowmsgsender = _allowed[msg_sender][spender] in + _allowed[msg_sender][spender] := allowmsgsender + addedValue; + emit Approval(msg_sender, spender, allowmsgsender); + true + + let decreaseApproval (_spender, _subtractedValue) = + pausable.whenNotPaused(); + let oldValue = _allowed[msg_sender][_spender] in + if (_subtractedValue > oldValue) then begin + _allowed[msg_sender][_spender] := 0u0; + let a = _allowed[msg_sender][_spender] in + emit Approval(msg_sender, _spender, a); + true + end else begin + _allowed[msg_sender][_spender] := oldValue - _subtractedValue; + let b = _allowed[msg_sender][_spender] in + emit Approval(msg_sender, _spender, b); + true + end +} + +layer TOKEN : [{ pausable : PausableInterface; ownable : OwnableInterface; lockable : LockableInterface}] { token : ERC20Interface } += { +token = Token +} + +layer CONTRACT = TOKEN @ LOCKABLE @ PAUSABLE @ OWNABLE + diff --git a/minic_unittests/contracts/token/token.ds b/minic_unittests/contracts/token/token.ds new file mode 100755 index 0000000..3b005e8 --- /dev/null +++ b/minic_unittests/contracts/token/token.ds @@ -0,0 +1,61 @@ +(* Loosely based on the "sample fixed supply token contract" from https://theethereum.wiki/w/index.php/ERC20_Token_Standard#Sample_Fixed_Supply_Token_Contract *) + +external [[except DataTypeOps]] with "Require token.Invariant.\n" + +const _totalSupply = 100000 + +object signature ERC20Interface = { + constructor : unit -> unit; + const totalSupply : unit -> int; + const balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object FixedSupplyToken : ERC20Interface { + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + balances[msg_sender] := 100000 + + let totalSupply () = + let bal0 = balances[address(0)] in + _totalSupply - bal0 + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + bal + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + true + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + true + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + true + +} + +layer signature FIXEDSUPPLYTOKENSig = { + fixedSupplyToken : ERC20Interface +} + +layer FIXEDSUPPLYTOKEN : [{}] FIXEDSUPPLYTOKENSig = { + fixedSupplyToken = FixedSupplyToken +} assert "Invariant.inv" diff --git a/minic_unittests/contracts/token/token/FunctionalCorrectness.v b/minic_unittests/contracts/token/token/FunctionalCorrectness.v new file mode 100755 index 0000000..dabf404 --- /dev/null +++ b/minic_unittests/contracts/token/token/FunctionalCorrectness.v @@ -0,0 +1,81 @@ +Require Import token.DataTypeOps. +Require Import token.LayerFIXEDSUPPLYTOKEN. +Require Import token.Invariant. + +Require Import DeepSpec.lib.Monad.StateMonadOption. +Require Import DeepSpec.lib.Monad.RunStateTInv. +Require Import lib.ArithInv. +Import DeepSpec.lib.Monad.Monad.MonadNotation. + +Require Import ZArith. +Require Import cclib.Maps. +Require Import cclib.Integers. + +Section WithMem. + +Import core.MemoryModel. + +(* +(* Todo: move this lemma to inv_arith. *) +Lemma cmpu_Cne_true : forall x y, Int256.cmpu Cne x y = true -> x<>y. +Proof. + intros. + unfold Int256.cmpu in H. + rewrite Bool.negb_true_iff in H. + apply Int256eq_false in H. + assumption. +Qed. *) + +(* We have now loaded the specification of the transfer method. *) +(* Print FixedSupplyToken_transfer_opt. *) + +Transparent balances_sum. + +Context {memModelOps : MemoryModelOps mem}. + +(* We can now prove that the transfer method does not create or destroy tokens. *) +Theorem transfer_constant_balances_sum : forall toA n d d' me b, + runStateT (FixedSupplyToken_transfer_opt toA n me) d = Some (b, d') + -> balances_sum d' = balances_sum d. +Proof. + intros. + Transparent FixedSupplyToken_transfer_opt. + unfold FixedSupplyToken_transfer_opt in H. + inv_runStateT. + subst. + inv_arith. + unfold balances_sum. + subst. + autorewrite with updates. + remember (FixedSupplyToken_balances m2) as m. + rewrite Int256Tree_Properties.sum_swap by congruence. + apply Int256Tree_Properties.constant_sum'. + + reflexivity. + + reflexivity. + + congruence. +Qed. + +(* And similar for the transferFrom method. (The proof is actually + identical.) *) +Theorem transferFrom_constant_balances_sum : forall fromA toA n d d' me b, + runStateT (FixedSupplyToken_transferFrom_opt fromA toA n me) d = Some (b, d') + -> balances_sum d' = balances_sum d. +Proof. + intros. + Transparent FixedSupplyToken_transferFrom_opt. + unfold FixedSupplyToken_transferFrom_opt in H. + inv_runStateT. + subst. + inv_arith. + unfold balances_sum. + subst. + autorewrite with updates. + remember (FixedSupplyToken_balances m2) as m. + rewrite Int256Tree_Properties.sum_swap by congruence. + apply Int256Tree_Properties.constant_sum'. + + reflexivity. + + reflexivity. + + congruence. +Qed. + +End WithMem. diff --git a/minic_unittests/contracts/token/token/Invariant.v b/minic_unittests/contracts/token/token/Invariant.v new file mode 100755 index 0000000..914975f --- /dev/null +++ b/minic_unittests/contracts/token/token/Invariant.v @@ -0,0 +1,25 @@ +Require Import ZArith. + +Require Import cclib.Maps. +Require Import token.DataTypeOps. + +(* This must match the constant defined in the token.ds file, it would be + better if the frontend generated this definition. *) +Definition _totalSupply : Z := 100000 %Z. + +Definition balances_sum (d : global_abstract_data_type) : Z := + Int256Tree_Properties.sum (FixedSupplyToken_balances d). + +Definition balances_nonnegative d := + forall i n, + Int256Tree.get i (FixedSupplyToken_balances d) = Some n + -> (0 <= n)%Z. + +(* Additionally we need that all the balance entries are non-negative, + but we will get that from the hashmap ft_cond. *) + +Definition inv d := + balances_nonnegative d /\ + balances_sum d = _totalSupply. + +Global Opaque balances_nonnegative balances_sum. (* to prevent it from being expanded when defining the VC. *) \ No newline at end of file diff --git a/minic_unittests/contracts/token/token/ManualRefinement.v b/minic_unittests/contracts/token/token/ManualRefinement.v new file mode 100755 index 0000000..f617932 --- /dev/null +++ b/minic_unittests/contracts/token/token/ManualRefinement.v @@ -0,0 +1,549 @@ +Require Import backend.phase.MiniC.Language. +Require Import backend.phase.MiniC.BigstepSemantics. +Require Import backend.phase.MiniC.Semantics. +Require Import backend.Values. +Require Import backend.MemoryModel. +Require Import core.MemoryModel. + +Section STMT_FUNC. + +Context `{HM : HyperMem}. + +(* -----------------stuff that should go in core.SynthesisStmt ------------------ *) + +Require Import cclib.Maps. +Require Import core.HyperType. +Require Import DeepSpec.core.SEnv. +Require Import DeepSpec.core.Syntax. + + (* todo: move this: *) + Definition empty_temp_env := PTree.empty val. + + Let env := empty_temp_env. + + (*Variable m0 : MemLow.*) + (* --------------------- actual refinement proof ------------------ *) + +Require Import DeepSpec.Runtime. +Require Import token.EdsgerIdents. +Require Import token.DataTypes. +Require Import token.DataTypeOps. +Require Import token.DataTypeProofs. +Require Import token.LayerFIXEDSUPPLYTOKEN. +Require Import token.Invariant. +Require Import token.ObjFixedSupplyTokenCodeProofs. +Require Import token.RefineFIXEDSUPPLYTOKEN. + + +Require Import lib.Monad.StateMonad. + +Context (ge : genv) (me : MachineModel.machine_env). + +(* Note, later we should fix the the (fst ml) by making the operational + semantics thread through an abstract datatype also. *) + +(* Context`{LayerSpec : LayerSpecClass}. (* TODO: this line will go away when we fix the type of HyperMem. *) *) +Set TypeClasses Debug. + +(*Context {memModelOps : MemoryModelOps storage_env}. *) +Context`{CTXT_prf : !Layer_FIXEDSUPPLYTOKEN_Context_prf}. + +Existing Instances RefineFIXEDSUPPLYTOKEN.memModelOps RefineFIXEDSUPPLYTOKEN.CTXT_prf FIXEDSUPPLYTOKEN_hypermem GlobalLayerSpec FIXEDSUPPLYTOKEN_overlay_spec FIXEDSUPPLYTOKEN_underlay_spec. + +Require Import lib.Monad.RunStateTInv. + + +(* This should go in a generated file somewhere, + probably ObjFixedSupplyTokenCode.v. *) +Instance FixedSupplyToken_balances_var_hyperltype : + HyperLType FixedSupplyToken_balances_var. +Proof. + constructor. + constructor. + - (* ltype_get_match *) + intros j d m mm dc; split. + + apply mm. + + apply mm. + - (* ltype_set_match *) + intros f j d m fc mm dc m' disjoint_eq same_d match_indirect. + constructor. + + destruct mm as [Hrelate _]. (* relate_AbData *) + rewrite same_d in *. + destruct Hrelate. + constructor; auto. + + constructor. (* match_AbData *) + split. + * exact match_indirect. + * exact fc. +Qed. + +(* This ad-hoc lemma will probably not be needed in the generic proofs. *) +Lemma match_indirect_basic: forall m ty i cv, + is_immediate_type ty -> + cval_match_indirect m ty i cv -> + cval_match (IdentExtMap.get i m) cv. +Proof. + intros m ty i cv H mm. + inversion mm; subst; try (now inversion H). + assumption. +Qed. + + +Theorem totalSupply_correctness : +(* forall wf ge M, MakeProgram.make_globalenv GetLowDataX (M, GetLowLayer) = Errors.OK ge -> *) + forall f d d', + runStateT (FixedSupplyToken_totalSupply_opt me) d = Some (f, d') -> + forall j ml (mcond : mem_match j d ml) lg, + synth_func_cond FixedSupplyToken_totalSupply FixedSupplyToken_totalSupply_wf me d -> + exists ml' lg' g' vres, + mem_match j d' ml' /\ + ht_rel (tp:=tint_Z32) f vres /\ + eval_funcall ge me (fst ml) lg + FixedSupplyToken_totalSupply_cfun (nil (*<-args*)) + (fst ml') lg' g' vres. +Proof. + intros f d d' Hrun j ml mcond lg Hvc. + assert (mcond_copy := mcond). + destruct mcond_copy as [? [[Hmatch_balances]]]. + Transparent FixedSupplyToken_totalSupply_opt. + unfold FixedSupplyToken_totalSupply_opt in Hrun. + inv_runStateT. + subst. + + unfold FixedSupplyToken_totalSupply_cfun. + +(* We first want to build a lens corresponding to the hashtable itself, + and then to one particular dereferenced location in it. + (LChash tint_Z32 + (LCvar FixedSupplyToken_balances_var) + (ECconst_int256 tint_U (Int256.repr 0) (Int256.repr 0))) +*) + pose (l_balance := inthash_ltype_pair FixedSupplyToken_balances_var (Int256.zero)). +g pose (balance0 := l_balance.(ltype_get) d'). + + assert (ht_ft_cond balance0 /\ + cval_match_indirect (fst ml) (unpair_ty tint_Z32) + (Ihash (Iident var_FixedSupplyToken_balances_ident) Int256.zero) + (ht_cval (balance0))) + as H_balance_get. + { + assert (balance0_hyperltype : HyperLType (inthash_ltype_pair FixedSupplyToken_balances_var Int256.zero)). + { eapply inthash_ltype. + eapply FixedSupplyToken_balances_var_hyperltype. + apply thash_int_HASH_Z_Z32_hash. + } + + destruct balance0_hyperltype as [balance0_hyperltype_dir]. + specialize (balance0_hyperltype_dir eq_refl). + destruct (ltype_get_match j d' ml mcond) as [balance0_fc balance0_match]. + { simpl. exact I. } + simpl in balance0_fc, balance0_match. + split; [exact balance0_fc | exact balance0_match]. + } + destruct H_balance_get as [balance0_fc balance0_match]. + + (* Next, for the arithmetic operation we want to prove something like + exists v, + (forall m, eval_expr ge env le m (synth_expr_expr tmp e) v) /\ + ht_rel (synth_expr_spec tmp e wf se) v + *) + assert (bin_cond : @Hbinary_cond Osub tint_Z32 tint_Z32 tint_Z32 + int_Z32_sub_impl _totalSupply balance0). + { simpl. + + cbv -[Int256.modulus zeq zle zlt Z.iter Z.le Z.lt Z.gt Z.ge Z.eqb Z.leb Z.ltb Z.geb Z.gtb Z.mul Z.div Z.modulo Z.add Z.sub Z.shiftl Z.shiftr Z.lxor Z.land Z.lor Int256.add Int256.sub Int256.mul Int256.modu Int256.divu Int256.cmpu Int256.not Int256.and Int256.or Int256.xor Int256.shl Int256.shru Ziteri Z.of_nat List.length + (*ret bind mzero get put gets guard modify runStateT evalStateT execStateT *) + is_true bool_dec ZMap.get ZMap.set Int256Tree.get Int256Tree.set hlist_hd + + balances update_balances] + in Hvc. + destruct Hvc as [[_ Hvc] _]. + specialize (Hvc balance0 d'). + destruct Hvc as [_ Hvc]. + { + reflexivity. + } + apply Hvc. + auto. + } + + apply match_indirect_basic in balance0_match; [|constructor]. + + + assert (fc_totalSupply : ht_ft_cond (tp:=tint_Z32) _totalSupply). + { + split; reflexivity. + } + assert (bin_correct := Hbinary_correct _ _ + (ht_cval (tp:=tint_Z32) _totalSupply) + (ht_cval (tp:=tint_Z32) balance0) + fc_totalSupply balance0_fc + bin_cond eq_refl eq_refl). + unfold ht_cval_some in bin_correct. + assert (bin_correct' : + Some (ht_cval (tp:=tint_Z32) (Hbinary (HyperBinaryImpl:=int_Z32_sub_impl) (tpl:=tint_Z32) _totalSupply balance0)) = + (cval_binary_operation Osub (ht_cval (tp:=tint_Z32) _totalSupply) + (unpair_ty tint_Z32) (ht_cval balance0) + (unpair_ty tint_Z32))). + { remember _totalSupply as _totalSupply'. + inversion bin_correct. + simpl. + f_equal. + subst. + exact H3. + } + clear bin_correct. + symmetry in bin_correct'. + + assert (Hmatch_totalSupply : cval_match (Vint (Int256.repr _totalSupply)) (ht_cval (tp:=tint_Z32) _totalSupply)). + { + apply CVMval. + } + eapply (cval_sem_binary_operation _ _ _ _ _ _ _ _ Hmatch_totalSupply balance0_match) in bin_correct'. + destruct bin_correct' as [v [bin_correct_step bin_correct_match]]. + + + destruct (TempModel.function_initial_temps_defined 12%positive tint ftype FixedSupplyToken_totalSupply_cfun (@PTree.empty _ : temp_env)) as [v12 Hv12] ;[simpl; auto | simpl;tauto | ]. + + destruct (TempModel.function_initial_temps_defined 10%positive tint ftype FixedSupplyToken_totalSupply_cfun (@PTree.empty _ : temp_env)) as [v10 Hv10] ;[simpl; auto | simpl;tauto | ]. + + + (* Yay, that's all the setup we need, now we can run the operational + semantics. *) + + eexists. + eexists. + eexists. + eexists. + split; [|split]. + - exact mcond. + - (* The return value. *) + exact bin_correct_match. + - eapply eval_funcall_internal. + + reflexivity. + + match goal with + [|- exec_stmt _ _ ?TEMPS _ _ _ _ _ _ _ _] => + set (the_temps := TEMPS) + end. + simpl. + + eapply exec_Sseq_1. + eapply exec_Sseq_1. + eapply exec_Sset. + { exact Hv12. } + eapply eval_Elvalue. + eapply eval_Ehashderef. + eapply eval_Evar. + eapply eval_Econst_int256. + + (* eapply exec_Sseq_1. *) + eapply exec_Sset. + { exact Hv10. } + eapply eval_Ebinop. + eapply eval_Econst_int256. + eapply eval_Etempvar. rewrite PTree.gss. reflexivity. + + exact bin_correct_step. + + eapply exec_Sreturn_some. + eapply eval_Etempvar. + rewrite PTree.gss. + reflexivity. +Qed. + +(* + +Theorem transfer_correctness : +(* forall wf ge M, MakeProgram.make_globalenv GetLowDataX (M, GetLowLayer) = Errors.OK ge -> *) + forall toA toA_val tokens tokens_val f d d', + (* These things come from "good_values" in SynthesisFunc.v *) + ht_ft_cond (tp:=tint_U) toA -> ht_rel (tp:=tint_U) toA toA_val -> + ht_ft_cond (tp:=tint_Z32) tokens -> ht_rel (tp:=tint_Z32) tokens tokens_val -> + runStateT (FixedSupplyToken_transfer_opt toA tokens me) d = Some (f, d') -> + forall j ml (mcond : mem_match j d ml) lg, + synth_func_cond FixedSupplyToken_transfer FixedSupplyToken_transfer_wf toA tokens me d -> + exists ml' lg' g' vres, + mem_match j d' ml' /\ + ht_rel (tp:=tint_bool) f vres /\ + eval_funcall ge me (fst ml) lg + FixedSupplyToken_transfer_cfun (toA_val :: tokens_val :: nil (*<-args*)) + (fst ml') lg' g' vres. +Proof. + intros toA toA_val tokens tokens_val f d d' + Hft_cond_toA Hrel_toA Hft_cond_tokens Hrel_tokens + Hrun j ml mcond lg Hvc. + assert (mcond_copy := mcond). + destruct mcond_copy as [? [[Hmatch_balances]]]. + Transparent FixedSupplyToken_transfer_opt. + unfold FixedSupplyToken_transfer_opt in Hrun. + remember (Hquery0 me) as fromA. + + assert (fromA_rel : ht_cval fromA = (CVval (MachineModel.me_query me (MachineModel.Qcall0 Hbuiltin0)))). + { + subst. + apply Hbuiltin_correct. + } + (* hide the HeqfromA equation from `subst`. *) + change (fromA = Hquery0 me) with ((fun x=>x) (fromA = Hquery0 me)) in HeqfromA. + + inv_runStateT. + destruct (Integers.Int256.cmpu Cne fromA toA)%Z eqn:?. + match goal with [H: runStateT (if ?X then _ else _) _ = _ |-_] => + destruct X eqn:? end. + inv_runStateT;subst. + (* There are three goals, corresponding to each path through the if-statements. + The proof for each one is similar. + *) + { + + unfold FixedSupplyToken_transfer_cfun. + rename x into fromA. + rename x4 into d. + + pose (l_balance_fromA := inthash_ltype_pair FixedSupplyToken_balances_var fromA). + pose (balance_fromA := l_balance_fromA.(ltype_get) d). + assert (ht_ft_cond balance_fromA /\ + cval_match_indirect (fst ml) (unpair_ty tint_Z32) + (Ihash (Iident var_FixedSupplyToken_balances_ident) fromA) + (ht_cval (balance_fromA))) + as H_balance_get. + { + assert (balance_fromA_hyperltype : HyperLType (inthash_ltype_pair FixedSupplyToken_balances_var fromA)). + { eapply inthash_ltype. + eapply FixedSupplyToken_balances_var_hyperltype. + apply thash_int_HASH_Z_Z32_hash. + } + + destruct balance_fromA_hyperltype as [balance_fromA_hyperltype_dir]. + specialize (balance_fromA_hyperltype_dir eq_refl). + destruct (ltype_get_match j d ml mcond) as [balance_fromA_fc balance_fromA_match]. + { simpl. exact I. } + simpl in balance_fromA_fc, balance_fromA_match. + split; [exact balance_fromA_fc | exact balance_fromA_match]. + } + destruct H_balance_get as [balance_fromA_fc balance_fromA_match]. + + +(* destruct (TempModel.function_initial_temps_defined2 ftype FixedSupplyToken_transfer_cfun (tokens_val :: toA_val :: nil) 11%positive tint + (PTree.set 15 (Vint Int256.zero) + (PTree.set 14 + (Vint Int256.zero) + (PTree.set 13 + (Vint Int256.zero) + (PTree.set 10 + (Vint Int256.zero) + (PTree.set 11 tokens_val (PTree.set 12 toA_val (@PTree.empty val)))))))) + as [v11 Hv11]; + [simpl; auto | reflexivity | ]. *) + + destruct (TempModel.function_initial_temps_defined2 ftype FixedSupplyToken_transfer_cfun (toA_val :: tokens_val::nil) 11%positive tint + (PTree.set 15 (Vint Int256.zero) + (PTree.set 14 + (Vint Int256.zero) + (PTree.set 13 + (Vint Int256.zero) + (PTree.set 10 + (Vint Int256.zero) + (PTree.set 11 toA_val (PTree.set 12 tokens_val (@PTree.empty val)))))))) + as [v11 Hv11]; + [simpl; auto | reflexivity | ]. + + (* Fixme: this should not actually be PTree.empty, it should be init_args. *) + destruct (TempModel.function_initial_temps_defined 13%positive tint ftype FixedSupplyToken_transfer_cfun (@PTree.empty _ : temp_env)) as [v13 Hv13]; [simpl; tauto | simpl;intuition;congruence | ]. + destruct (TempModel.function_initial_temps_defined 14%positive tint ftype FixedSupplyToken_transfer_cfun (@PTree.empty _ : temp_env)) as [v14 Hv14]; [simpl; tauto | simpl;intuition;congruence | ]. + destruct (TempModel.function_initial_temps_defined 15%positive tint ftype FixedSupplyToken_transfer_cfun (@PTree.empty _ : temp_env)) as [v15 Hv15]; [simpl; tauto | simpl;intuition;congruence | ]. + + assert (Hft_cond_fromA : ht_ft_cond (tp:=tint_U) fromA). + { + split; reflexivity. + } + assert (bin_correct_ne := Hbinary_correct (op := One) _ _ + (ht_cval (tp:=tint_U) fromA) + (ht_cval (tp:=tint_U) toA) + Hft_cond_fromA Hft_cond_toA + I eq_refl eq_refl). + unfold ht_cval_some in bin_correct_ne. + assert (bin_correct_ne' : + Some (ht_cval (tp:=tint_bool) (Hbinary (HyperBinaryImpl:=int_U_ne_impl) (tpl:=tint_U) fromA toA)) = + (cval_binary_operation One (ht_cval (tp:=tint_U) fromA) + (unpair_ty tint_U) (ht_cval toA) + (unpair_ty tint_bool))). + { inversion bin_correct_ne. + simpl. + f_equal. + } + clear bin_correct_ne. + symmetry in bin_correct_ne'. + + unfold tint_U in bin_correct_ne'. + unfold addr in bin_correct_ne'. + rewrite fromA_rel in bin_correct_ne'. + + assert (Hmatch_fromA : cval_match (Vint fromA) (CVval (MachineModel.me_query me (MachineModel.Qcall0 Hbuiltin0)))). + { + rewrite <- fromA_rel. + apply CVMval. + } + + eapply (cval_sem_binary_operation _ _ _ _ _ _ _ _ Hmatch_fromA Hrel_toA) in bin_correct_ne'. + destruct bin_correct_ne' as [v_ne [bin_correct_ne_step bin_correct_ne_match]]. + + assert (bin_cond_ge : @Hbinary_cond Oge tint_Z32 tint_Z32 tint_bool + int_Z32_ge_impl balance_fromA tokens) + by exact I. + + apply match_indirect_basic in balance_fromA_match; [|constructor]. + assert (bin_correct_ge := Hbinary_correct _ _ + (ht_cval (tp:=tint_Z32) balance_fromA) + (ht_cval (tp:=tint_Z32) tokens) + balance_fromA_fc + Hft_cond_tokens + bin_cond_ge eq_refl eq_refl). + unfold ht_cval_some in bin_correct_ge. + assert (bin_correct_ge' : + Some (ht_cval (tp:=tint_bool) (Hbinary (HyperBinaryImpl:=int_Z32_ge_impl) (tpl:=tint_Z32) balance_fromA tokens )) = + (cval_binary_operation Oge + (ht_cval balance_fromA) (unpair_ty tint_Z32) + (ht_cval (tp:=tint_Z32) tokens) + (unpair_ty tint_bool))). + { + inversion bin_correct_ge. + simpl. + f_equal. + subst. + exact H3. + } + clear bin_correct_ge. + symmetry in bin_correct_ge'. + apply (cval_sem_binary_operation _ _ _ _ _ _ _ _ balance_fromA_match Hrel_tokens) in bin_correct_ge'. + destruct bin_correct_ge' as [v_ge [bin_correct_ge_step bin_correct_ge_match]]. + + pose (ne_result := Hbinary (HyperBinaryImpl:=int_U_ne_impl) fromA toA). + pose (ge_result := Hbinary (HyperBinaryImpl:=int_Z32_ge_impl) balance_fromA tokens). + + assert (bin_cond_and : @Hbinary_cond Oand tint_bool tint_bool tint_bool + int_bool_and_impl ne_result ge_result) + by exact I. + + assert (bin_correct_and := Hbinary_correct (HyperBinaryImpl0:=int_bool_and_impl) + _ _ + (ht_cval (tp:=tint_bool) ne_result) + (ht_cval (tp:=tint_bool) ge_result) + (Hbinary_returns _ _ Hft_cond_fromA Hft_cond_toA I) + (Hbinary_returns _ _ balance_fromA_fc Hft_cond_tokens bin_cond_ge) + bin_cond_and eq_refl eq_refl). + unfold ht_cval_some in bin_correct_and. + assert (bin_correct_and' : + Some (ht_cval (tp:=tint_bool) (Hbinary (HyperBinaryImpl:=int_bool_and_impl) (tpl:=tint_bool) ne_result ge_result )) = + (cval_binary_operation Oand + (ht_cval ne_result) (unpair_ty tint_bool) + (ht_cval (tp:=tint_bool) ge_result) + (unpair_ty tint_bool))). + { + inversion bin_correct_and. + simpl. + f_equal. + subst. + exact H3. + } + clear bin_correct_and. + symmetry in bin_correct_and'. + apply (cval_sem_binary_operation _ _ _ _ _ _ _ _ bin_correct_ne_match bin_correct_ge_match) in bin_correct_and'. + destruct bin_correct_and' as [v_and [bin_correct_and_step bin_correct_and_match]]. + + + eexists. + eexists. + eexists. + eexists. + split; [|split]. + + - admit. (* exact mcond *) + - (* the return value *) + unfold ht_rel. + simpl. + apply CVMval. + - eapply eval_funcall_internal. + + reflexivity. + + match goal with + [|- exec_stmt _ _ ?TEMPS _ _ _ _ _ _ _ _] => + set (the_temps := TEMPS) + end. + simpl. + + eapply exec_Sseq_1. + eapply exec_Sseq_1. + eapply exec_Sset. + { exact Hv13. } + eapply eval_Ecall0. + reflexivity. + eapply exec_Sseq_1. + eapply exec_Sset. + { rewrite PTree.gso. + exact Hv14. + congruence. + } + eapply eval_Elvalue. + eapply eval_Ehashderef. + eapply eval_Evar. + eapply eval_Etempvar. + { unfold ht_cval in fromA_rel. + simpl in fromA_rel. + inversion fromA_rel as [fromA_rel']. + symmetry. + rewrite PTree.gss. + f_equal. + exact fromA_rel'. + } + eapply exec_Sseq_1. + eapply exec_Sset. + { repeat rewrite PTree.gso by congruence. + exact Hv15. + } + + eapply eval_Elvalue. + eapply eval_Ehashderef. + eapply eval_Evar. + eapply eval_Etempvar. + { + repeat rewrite PTree.gso by congruence. + unfold the_temps. + (* stuck here exact Hv11. *) admit. + } + + (* We are in the "then" branch of the if statement. *) + eapply exec_Sifthenelse. + { eapply eval_Ebinop. (* and *) + - eapply eval_Ebinop. (* <> *) + + eapply eval_Etempvar. + rewrite !PTree.gso by congruence. + rewrite PTree.gss. + reflexivity. + + eapply eval_Etempvar. + simpl in the_temps. + unfold the_temps. + simpl. + reflexivity. + + match goal with [|- sem_binary_operation One ?X _ _ _ = _] => + replace X with (Vint fromA) + by (inversion fromA_rel as [fromA_rel_eq]; rewrite fromA_rel_eq; reflexivity) + end. + exact bin_correct_ne_step. + - eapply eval_Ebinop. (* >= *) + + eapply eval_Etempvar. + rewrite !PTree.gso by congruence. + rewrite PTree.gss. + reflexivity. + + eapply eval_Etempvar. + simpl in the_temps. + unfold the_temps. + simpl. + reflexivity. + + exact bin_correct_ge_step. + - exact bin_correct_and_step. +*) + + +End STMT_FUNC. + + diff --git a/minic_unittests/contracts/token/token/ObjFixedSupplyTokenCodeProofs.v b/minic_unittests/contracts/token/token/ObjFixedSupplyTokenCodeProofs.v new file mode 100644 index 0000000..93b6963 --- /dev/null +++ b/minic_unittests/contracts/token/token/ObjFixedSupplyTokenCodeProofs.v @@ -0,0 +1,249 @@ +(* Skeleton by Edgser for token.ds *) +Require Import BinPos. +Require Import DeepSpec.Runtime. +Require Import DeepSpec.core.SynthesisFunc. +Require Import token.EdsgerIdents. +Require Import token.DataTypes. +Require Import token.DataTypeOps. +Require Import token.DataTypeProofs. +(*Require Import liblayers.compcertx.MakeProgram. +Require Import liblayers.compcertx.MemWithData.*) + +Require Import token.LayerFIXEDSUPPLYTOKEN. +Require Import lib.ArithInv. +Require Import lib.Monad.RunStateTInv. +Require Import token.Invariant. +Transparent balances_sum balances_nonnegative. + +Section EdsgerGen. + +Existing Instance GlobalLayerSpec. + +Existing Instances FIXEDSUPPLYTOKEN_overlay_spec + FIXEDSUPPLYTOKEN_data_ops + FIXEDSUPPLYTOKEN_data. + +Context {memModelOps : MemoryModelOps mem}. + +Lemma FixedSupplyToken_constructor_vc me d : + high_level_invariant d -> + synth_func_cond FixedSupplyToken_constructor FixedSupplyToken_constructor_wf + me d. +Proof. + intros. + apply FixedSupplyToken_constructor_cond_eq; auto. + unfold FixedSupplyToken_constructor_cond. + intuition. +Qed. + +Lemma FixedSupplyToken_constructor_oblg me d : + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_constructor FixedSupplyToken_constructor_wf + me d. +Proof. + intros. + apply FixedSupplyToken_constructor_obligation_eq; auto. + unfold FixedSupplyToken_constructor_obligation. + intuition. +Qed. + + +Lemma FixedSupplyToken_totalSupply_vc me d : + high_level_invariant d -> + synth_func_cond FixedSupplyToken_totalSupply FixedSupplyToken_totalSupply_wf + me d. +Proof. + intros. + apply FixedSupplyToken_totalSupply_cond_eq; auto. + unfold FixedSupplyToken_totalSupply_cond. + intuition. + inv_runStateT. + subst. + apply Zle_ge. + apply Int256Tree_Properties.sum_bound1. + + intros. + apply Zle_ge. + unfold balances_nonnegative in H1. + eauto. + + unfold balances_sum in H2. + omega. +Qed. + +Lemma FixedSupplyToken_totalSupply_oblg me d : + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_totalSupply FixedSupplyToken_totalSupply_wf + me d. +Proof. + intros. + apply FixedSupplyToken_totalSupply_obligation_eq; auto. + unfold FixedSupplyToken_totalSupply_obligation. + intuition. +Qed. + +Lemma FixedSupplyToken_balanceOf_vc a0 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + high_level_invariant d -> + synth_func_cond FixedSupplyToken_balanceOf FixedSupplyToken_balanceOf_wf + a0 me d. +Proof. + intros. + apply FixedSupplyToken_balanceOf_cond_eq; auto. + unfold FixedSupplyToken_balanceOf_cond. + intuition. +Qed. + +Lemma FixedSupplyToken_balanceOf_oblg a0 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_balanceOf FixedSupplyToken_balanceOf_wf + a0 me d. +Proof. + intros. + apply FixedSupplyToken_balanceOf_obligation_eq; auto. + unfold FixedSupplyToken_balanceOf_obligation. + intuition. +Qed. + + +Lemma FixedSupplyToken_transfer_vc a0 a1 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + high_level_invariant d -> + synth_func_cond FixedSupplyToken_transfer FixedSupplyToken_transfer_wf + a0 a1 me d. +Proof. + intros. + apply FixedSupplyToken_transfer_cond_eq; auto. + unfold FixedSupplyToken_transfer_cond. + intuition. + + + inv_runStateT. + subst. + inv_arith. + auto. + + inv_runStateT. + subst. + inv_arith. + assert (100000 < Int256.modulus) by reflexivity. + + assert (Hbound: Int256Tree_Properties.sum (FixedSupplyToken_balances g0) <= 100000). + { unfold balances_sum in H11. + omega. + } + + assert (Hnonnegative : forall k v, Int256Tree.get k (FixedSupplyToken_balances g0) = Some v -> v >= 0). + { + unfold balances_nonnegative in H6. + intros k v Hlookup. + specialize (H6 k v Hlookup). + omega. + } + + assert (l := Int256Tree_Properties.sum_bound2 H8 Hnonnegative Hbound). + apply Zle_lt_trans with (Int256Tree.get_default 0 a0 (FixedSupplyToken_balances g0) + Int256Tree.get_default 0 (MachineModel.me_caller me) (FixedSupplyToken_balances g0)). + omega. + apply Zle_lt_trans with 100000. + rewrite Zplus_comm. + exact l. + omega. +Qed. + +Lemma FixedSupplyToken_transfer_oblg a0 a1 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_transfer FixedSupplyToken_transfer_wf + a0 a1 me d. +Proof. + intros. + apply FixedSupplyToken_transfer_obligation_eq; auto. + unfold FixedSupplyToken_transfer_obligation. + intuition. +Qed. + +Lemma FixedSupplyToken_approve_vc a0 a1 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + high_level_invariant d -> + synth_func_cond FixedSupplyToken_approve FixedSupplyToken_approve_wf + a0 a1 me d. +Proof. + intros. + apply FixedSupplyToken_approve_cond_eq; auto. + unfold FixedSupplyToken_approve_cond. + intuition. +Qed. + +Lemma FixedSupplyToken_approve_oblg a0 a1 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_approve FixedSupplyToken_approve_wf + a0 a1 me d. +Proof. + intros. + apply FixedSupplyToken_approve_obligation_eq; auto. + unfold FixedSupplyToken_approve_obligation. + intuition. +Qed. + +Lemma FixedSupplyToken_transferFrom_vc a0 a1 a2 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + ht_ft_cond a2 -> ht_valid_ft_cond a2 -> + high_level_invariant d -> + synth_func_cond FixedSupplyToken_transferFrom FixedSupplyToken_transferFrom_wf + a0 a1 a2 me d. +Proof. + intros. + apply FixedSupplyToken_transferFrom_cond_eq; auto. + unfold FixedSupplyToken_transferFrom_cond. + intuition. + + inv_runStateT. + subst. + inv_arith. + auto. + + inv_runStateT. + subst. + unfold Int256.cmpu in *. + inv_arith. + clear H17. (* temp. *) + assert (100000 < Int256.modulus) by reflexivity. + + assert (Hbound: Int256Tree_Properties.sum (FixedSupplyToken_balances g0) <= 100000). + { unfold balances_sum in H15. + omega. + } + + assert (Hnonnegative : forall k v, Int256Tree.get k (FixedSupplyToken_balances g0) = Some v -> v >= 0). + { + unfold balances_nonnegative in H10. + intros k v Hlookup. + specialize (H10 k v Hlookup). + omega. + } + assert (l := Int256Tree_Properties.sum_bound2 H12 Hnonnegative Hbound). + + + apply Zle_lt_trans with (Int256Tree.get_default 0 a0 (FixedSupplyToken_balances g0) + Int256Tree.get_default 0 a1 (FixedSupplyToken_balances g0)). + omega. + apply Zle_lt_trans with 100000. + exact l. + omega. +Qed. + +Lemma FixedSupplyToken_transferFrom_oblg a0 a1 a2 me d : + ht_ft_cond a0 -> ht_valid_ft_cond a0 -> + ht_ft_cond a1 -> ht_valid_ft_cond a1 -> + ht_ft_cond a2 -> ht_valid_ft_cond a2 -> + high_level_invariant d -> + synth_func_obligation FixedSupplyToken_transferFrom FixedSupplyToken_transferFrom_wf + a0 a1 a2 me d. +Proof. + intros. + apply FixedSupplyToken_transferFrom_obligation_eq; auto. + unfold FixedSupplyToken_transferFrom_obligation. + intuition. +Qed. + +End EdsgerGen. diff --git a/minic_unittests/contracts/token/token/RefineFIXEDSUPPLYTOKEN.v b/minic_unittests/contracts/token/token/RefineFIXEDSUPPLYTOKEN.v new file mode 100755 index 0000000..bf7e8f4 --- /dev/null +++ b/minic_unittests/contracts/token/token/RefineFIXEDSUPPLYTOKEN.v @@ -0,0 +1,56 @@ +(* This should be automatically generated, but I put in a manually written version for now. *) + +Require Import BinNums. +Require Import DeepSpec.core.HyperTypeInst. +Require Import DeepSpec.Runtime. +Require Import DeepSpec.Linking. +Require Import token.EdsgerIdents. +Require Import token.DataTypes. +Require Import token.DataTypeOps. +Require Import token.DataTypeProofs. + +Require Import token.LayerFIXEDSUPPLYTOKEN. + + +(* +Context {mem}`{Hmem: Mem.MemoryModel mem}. +Context`{Hmwd: UseMemWithData mem}. +Context`{make_program_ops: !MakeProgramOps Clight.function Ctypes.type Clight.fundef Ctypes.type}. +Context`{Hmake_program: !MakeProgram Clight.function Ctypes.type Clight.fundef Ctypes.type}. +Context`{global_abdata : !GlobalAbData init_global_abstract_data global_low_level_invariant}. + +Context {memModelOps : MemoryModelOps mem}. + *) + +Existing Instances GlobalLayerSpec FIXEDSUPPLYTOKEN_overlay_spec FIXEDSUPPLYTOKEN_underlay_spec. + +Context`{CTXT_prf : !Layer_FIXEDSUPPLYTOKEN_Context_prf}. + +(* +Record relate_RData (j : meminj) (habd : GetHighData) (labd : GetLowData) : Prop + := mkrelate_RData { +}.*) + +Record match_RData (habd : GetHighData) (m : mem) (j : MemoryModel.meminj) : Prop + := MATCH_RDATA { + balances_ma : variable_match FixedSupplyToken_balances_var habd m +}. + +Local Hint Resolve MATCH_RDATA. + +Global Instance rel_ops: CompatRelOps GetHighDataX GetLowDataX := +{ + relate_AbData f d1 d2 := True; + match_AbData d1 m f := match_RData d1 m f; + new_glbl := var_FixedSupplyToken_balances_ident :: nil +}. + + +Global Instance rel_prf: CompatRel GetHighDataX GetLowDataX. + + +Existing Instance GlobalLayerSpec. + +Instance FIXEDSUPPLYTOKEN_hypermem : MemoryModel.HyperMem + := { Hcompatrel := {| crel_prf := rel_prf |} }. + diff --git a/minic_unittests/contracts/token/token/_CoqProject b/minic_unittests/contracts/token/token/_CoqProject new file mode 100755 index 0000000..9ddcc27 --- /dev/null +++ b/minic_unittests/contracts/token/token/_CoqProject @@ -0,0 +1,11 @@ +-R ../../.. DeepSpec +-R . token +EdsgerIdents.v +DataTypeOps.v +DataTypeProofs.v +DataTypes.v +LayerFIXEDSUPPLYTOKEN.v +RefineFIXEDSUPPLYTOKEN.v +ObjFixedSupplyTokenCodeProofs.v +FunctionalCorrectness.v +Invariant.v \ No newline at end of file diff --git a/minic_unittests/contracts/token/token/prf.v b/minic_unittests/contracts/token/token/prf.v new file mode 100644 index 0000000..4b09193 --- /dev/null +++ b/minic_unittests/contracts/token/token/prf.v @@ -0,0 +1,281 @@ +Require Import token.DataTypes. +Require Import token.DataTypeOps. +Require Import lib.Monad.StateMonadOption. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import ZArith. +Require Import core.HyperTypeInst. +Require Import backend.MachineModel. +Require Import token.LayerFIXEDSUPPLYTOKEN. + +Definition state := global_abstract_data_type. + +Definition init_state := init_global_abstract_data. + +Definition wei := Z. + +Definition addr := int256. + +Definition blocknumber := int256. + +Existing Instance GlobalLayerSpec. + +Section step. + + Context (contract_address : addr). + + Section mstep. + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Definition make_machine_env (caller: addr) + : machine_env state + := {| me_address := contract_address; + me_origin := caller; + me_caller := caller; (* need update after every control-flow transfer *) + me_callvalue := Int256.repr (0); + me_coinbase := coinbase; + me_timestamp := timestamp; + me_number := number; + me_balance := balance; + me_blockhash := blockhash; + (* not implemented *) + me_transfer _ _ _ _ _ := False; + me_callmethod _ _ _ _ _ _ _ _ _ _ := False; + me_log _ _ _ := prev_contract_state; + |}. + + Import MonadNotation. + + Definition lif {A:Type} + (caller : addr) + (cmd : machine_env state -> osT global_abstract_data_type A) + : osT state A := + st <- get;; + let me := make_machine_env caller in + match runStateT (cmd me) st with + | None => mzero + | Some (v, st') => put st' ;; ret v + end. + + (* osT state int256 = + state transformer with option monad, state is state, and value is int256 *) + Print osT. (* = fun D : Type => stateT D option *) + Print stateT. (* (S : Type) (m : Type -> Type) (t : Type) : Type := mkStateT + { runStateT : S -> m (t * S)%type } *) + + Print runStateT. (* takes monad transformer, state, returns m (t * S) *) + + (* How the state of the system changes in a single method call made by player p. *) + Inductive mstep u (st st' : state) : Prop := + | FixedSupplyToken_constructor_step : forall r , runStateT (FixedSupplyToken_constructor_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_totalSupply_step : forall r , runStateT (FixedSupplyToken_totalSupply_opt (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_balanceOf_step : forall r tokenOwner, runStateT (FixedSupplyToken_balanceOf_opt tokenOwner (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_transfer_step : forall r toA tokens, runStateT (FixedSupplyToken_transfer_opt toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_approve_step : forall r spender tokens, runStateT (FixedSupplyToken_approve_opt spender tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' + | FixedSupplyToken_transferFrom_step : forall r fromA toA tokens, runStateT (FixedSupplyToken_transferFrom_opt fromA toA tokens (make_machine_env u)) st = Some (r, st') -> mstep u st st' +. + + (* We can compute a new block by letting the players call the contract + in some arbitrary order. *) + Inductive multi_mstep : state -> state -> Prop := + | multi_mstep_reflexive : forall (st : state), multi_mstep st st + | multi_mstep_transitive : forall (st st' st'' : state) u, + multi_mstep st st' -> mstep u st' st'' -> multi_mstep st st''. + + + (* A block is sufficiently synchronous if every player got a chance to submit a + transaction to to. *) + (* TODO, I think this definition is fine for now, but it seems little too clever, + should probably re-state this using some straightforward tracking of the states + we pass through. *) + Definition multi_mstep_synchronous st1 st2 := + forall u, exists st st', + multi_mstep st1 st /\ mstep u st st' /\ multi_mstep st' st2. + + (* Here are a bunch of induction principles inspired by linear temporal logic. *) + + (* Prove that some property P holds "globally", i.e. for each state along a + path. + You can also specify a property Pprev which is known to hold for the + prev_contract_state. If not needed, it can just be True. + *) + Lemma multi_mstep_Global_ind : forall (Pprev P : state -> Prop), + (forall u st st', Pprev prev_contract_state -> P st -> mstep u st st' -> P st') -> + Pprev prev_contract_state -> forall st st', P st -> multi_mstep st st' -> P st'. + Proof. + induction 4; eauto. + Qed. + + (* + (* Prove that P holds "until" Q along a path. + "Until" is a liveness assertion, so we need the synchronicity assumption. *) + Lemma multi_mstep_Until_ind : forall (Pprev P Q : state -> Prop), + (forall p st st', Pprev prev_contract_state -> P st -> In p players + -> mstep p st st' -> (P st' \/ Q st')) -> + Pprev prev_contract_state -> + forall st, + P st -> exists st', multi_mstep st st' -> (P st' \/ Q st'). + Proof. + induction 4; eauto *) + + End mstep. + + + Definition Int256_incr x := Int256.add x Int256.one. + + Inductive bstep (n : blocknumber) : state -> state -> Prop := + | bstep_step : forall coinbase timestamp balance blockhash st st', + multi_mstep coinbase timestamp n balance blockhash st st st' -> + bstep n st st'. + + Inductive multi_bstep : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_reflexive : forall n (st : state), multi_bstep n st n st + | multi_bstep_transitive : forall n n' (st st' st'' : state), + multi_bstep n st n' st' -> bstep (Int256_incr n') st' st'' -> multi_bstep n st (Int256_incr n') st''. + + (* multi_bstep is the step relation without any synchronicity assumption. + This is sufficient to prove some safety properties, but for most interesting + theorems we instead need to use this synchronous version: *) + + Inductive bstep_synch (n : blocknumber) : state -> state -> Prop := + | bstep_synch_step : forall coinbase timestamp balance blockhash st st', + multi_mstep_synchronous coinbase timestamp n balance blockhash st st st' -> + bstep_synch n st st'. + + Inductive multi_bstep_synch : blocknumber -> state -> blocknumber -> state -> Prop := + | multi_bstep_synch_reflexive : forall n (st : state), multi_bstep_synch n st n st + | multi_bstep_synch_transitive : forall n n' (st st' st'' : state), + multi_bstep_synch n st n' st' -> bstep_synch (Int256_incr n') st' st'' -> multi_bstep_synch n st (Int256_incr n') st''. + + Lemma multi_bstep_Global_ind : forall (P : state -> Prop), + (forall coinbase timestamp number balance blockhash prev_block u st st', + P prev_block + -> P st + -> mstep coinbase timestamp number balance blockhash prev_block u st st' + -> P st') + -> forall n st n' st', + P st -> multi_bstep n st n' st' -> P st'. + Proof. + induction 3. + - auto. + - inversion H2; subst. + eapply multi_mstep_Global_ind with (st:=st') (prev_contract_state := st'). + + intros. + refine (H _ _ _ _ _ _ _ _ _ _ _ H6); auto. + + apply IHmulti_bstep; assumption. + + apply IHmulti_bstep; assumption. + + exact H3. + Qed. + End step. + + Section DeepSEAGenericProof. + + Lemma Int256Tree_reduce : forall (i: int256) (v: Z) (t: Int256Tree.t Z), Int256Tree.get_default 0%Z i (Int256Tree.set i v t) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gss . + reflexivity. + Qed. + + Lemma Int256Tree_mreduce : forall (i j : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set i v t)) = v. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gss. + reflexivity. + exact H. + Qed. + + Lemma Int256Tree_mireduce : forall (i j k : int256) (v v': Z) (t: Int256Tree.t Z), + i <> j -> + i <> k -> + j <> k -> + Int256Tree.get_default 0%Z i (Int256Tree.set j v' (Int256Tree.set k v t)) = + Int256Tree.get_default 0%Z i t. + Proof. + intros. + unfold Int256Tree.get_default. + rewrite Int256Tree.gso. + rewrite Int256Tree.gso. + reflexivity. + exact H0. + exact H. + Qed. + + Lemma add_sub_inv : forall (i j : Z32), (i + j - i)%Z = j. + Proof. + intros. + omega. + Qed. + End DeepSEAGenericProof. + + Section Proof. + Context (* (strategies : strategy_tuple) *) + (initial_balances : addr -> Z) + (contract_address : int256). + + Context (init_bt init_rt : int256) + (init_coinbase : int256) + (init_timestamp : int256) + (init_number : int256) + (init_blockhash : int256 -> int256) + (pre_init_state init_state : state). + + (* These are the parameters which are constant within a given block. *) + Context (coinbase : int256) + (timestamp : int256) + (number : int256) + (balance : int256 -> int256) + (blockhash : int256 -> int256) + (prev_contract_state : state). + + Require Import lib.Monad.RunStateTInv. + Require Import lib.ArithInv. + + Definition make_machine_env_wrapped prev_st user := + make_machine_env contract_address coinbase timestamp number balance blockhash prev_st user. + + Lemma make_machine_env_caller_eq : forall st caller, me_caller (make_machine_env_wrapped st caller) = caller. + Proof. auto. Qed. + + Lemma make_machine_env_address_eq : forall st caller, me_address (make_machine_env_wrapped st caller) = contract_address. + Proof. auto. Qed. + +Transparent FixedSupplyToken_constructor_opt. +Transparent FixedSupplyToken_totalSupply_opt. +Transparent FixedSupplyToken_balanceOf_opt. +Transparent FixedSupplyToken_transfer_opt. +Transparent FixedSupplyToken_approve_opt. +Transparent FixedSupplyToken_transferFrom_opt. + +Ltac rds := +unfold FixedSupplyToken_constructor_opt in *; +unfold FixedSupplyToken_totalSupply_opt in *; +unfold FixedSupplyToken_balanceOf_opt in *; +unfold FixedSupplyToken_transfer_opt in *; +unfold FixedSupplyToken_approve_opt in *; +unfold FixedSupplyToken_transferFrom_opt in *; +inv_runStateT; +subst; +inv_arith; +simpl; +try rewrite make_machine_env_caller_eq in *; +try rewrite make_machine_env_address_eq in *; +try rewrite Int256Tree_reduce in *; +try rewrite Int256Tree_mreduce in *; +try rewrite Int256Tree_mireduce in *; +auto. + +Theorem sample : forall n, n = n. + +End Proof. diff --git a/minic_unittests/contracts/token_ant/token.js b/minic_unittests/contracts/token_ant/token.js new file mode 100644 index 0000000..0f741f4 --- /dev/null +++ b/minic_unittests/contracts/token_ant/token.js @@ -0,0 +1,145 @@ +// node token.js token.bytecode +const Chain = require("@alipay/mychain/index.node") // download from https://tech.antfin.com/docs/2/107128 +const fs = require("fs") + +const accountKey = fs.readFileSync("", { encoding: "utf8" }) // read from customized .pem password. eg: "./certs/password.pem" +const accountPassword = "" //need to replace to customized .pem file's password + +const keyInfo = Chain.utils.getKeyInfo(accountKey, accountPassword) + +const passphrase = "" //need to replace to custom client.key's password +//setting +let opt = { + host: '47.103.163.48', //get from your free ant blockChain + port: 18130, //port number + timeout: 30000, + cert: fs.readFileSync("./certs/client.crt", { encoding: "utf8" }), // client.crt + ca: fs.readFileSync("./certs/ca.crt", { encoding: "utf8" }), // ca.crt + key: fs.readFileSync("./certs/client.key", { encoding: "utf8" }), // client key + userPublicKey: keyInfo.publicKey, + userPrivateKey: keyInfo.privateKey, + userRecoverPublicKey: keyInfo.publicKey, + userRecoverPrivateKey: keyInfo.privateKey, + passphrase: passphrase +} + +// initialize a connection +const chain = Chain(opt) + +/////////////this code is used for create a new account//////////////////////////////// +// /////////////// +// const newKey = Chain.utils.generateECKey(); /////////////// +// console.log('newKey priKey:', newKey.privateKey.toString('hex')) /////////////// +// console.log('newKey pubKey:', newKey.publicKey.toString('hex')) /////////////// +// /////////////// +// const newAccountName = 'Tester001' + Date.now() /////////////// +// chain.ctr.CreateAccount({ /////////////// +// from: 'Tester001', /////////////// +// to: newAccountName, /////////////// +// data: { /////////////// +// recover_key: '0x'+ newKey.publicKey.toString('hex'), /////////////// +// auth_key: '0x'+ newKey.publicKey.toString('hex'), /////////////// +// auth_weight: 100 /////////////// +// } /////////////// +// }, (err, data) => { /////////////// +// console.log('New Account:', data) /////////////// +// /////////////// +// opt.userPrivateKey = '0x'+ newKey.privateKey.toString('hex') /////////////// +// opt.userPublicKey = '0x'+ newKey.publicKey.toString('hex') /////////////// +// opt.userRecoverPrivateKey = '0x'+ newKey.privateKey.toString('hex')/////////////// +// opt.userRecoverPublicKey = '0x'+ newKey.publicKey.toString('hex') /////////////// +// /////////////// +// // chain.setUserKey(opt) /////////////// +// // chain.setUserRecoverKey(opt) /////////////// +// /////////////// +// }) /////////////// +// chain.ctr.QueryAccount({ /////////////// +// from: 'Tester001' /////////////// +// }, (err, data) => { /////////////// +// console.log('QueryAccount Tester001:', data) /////////////// +// }) /////////////// +// eg: /////////////// +// User name : Tester001 /////////////// +// /////////////// +/////////////////////////////////////////////////////////////////////////////////////// + +// ./dsc_antchain contracts/token/token.ds abi +const abi = [ {"type":"constructor", + "name":"constructor", + "inputs":[], + "outputs":[], + "payable":false, + "constant":false, + "stateMutability":"nonpayable"}, + {"type":"function", + "name":"totalSupply", + "inputs":[], + "outputs":[{"name":"", "type":"uint256"}], + "payable":false, + "constant":true, + "stateMutability":"view"}, + {"type":"function", + "name":"balanceOf", + "inputs":[{"name":"tokenOwner", "type":"identity"}], + "outputs":[{"name":"", "type":"uint256"}], + "payable":false, + "constant":true, + "stateMutability":"view"}, + {"type":"function", + "name":"transfer", + "inputs":[{"name":"toA", "type":"identity"},{"name":"tokens", "type":"uint256"}], + "outputs":[{"name":"", "type":"bool"}], + "payable":true, + "constant":false, + "stateMutability":"payable"}]; + +const bytecode = fs.readFileSync(process.argv[2]).toString().replace(/\n|\t|\r| /g, ""); + +// add date to make sure there is no duplicate hash +const contractName = 'contract'+Date.now() +// initalize a contract example +let myContract = chain.ctr.contract(contractName, abi) + +myContract.new(bytecode, { + from: '', // from your user name eg: Tester001 + parameters: [] +}, (err, contract, data) => { + console.log('contract deploy result:', data) + myContract.totalSupply({ + from: '' + }, (err, output, data) => { + // console.log("contract call error: ", err) + // console.log('contract call data: :', data) + console.log('contract call output:', output) + // replace fake identity by a real identity, example: 0xc60a9d48105950a0cca07a4c6320b98c303ad42d694a634529e8e1a0a16fcdb5 + myContract.balanceOf(('fake identity') , { + from: '' + }, (err, output, data) => { + // console.log("contract call error: ", err) + // console.log('contract call data: ', data) + console.log('contract call output:', output) + //100 is just an example, change to the number you want to send. + myContract.transfer('fake identity', 100 , { + from: '' + }, (err, output, data) => { + // console.log("contract call error: ", err) + // console.log('contract call data transfer: ', data) + console.log('contract call output transfer: ', output) + myContract.balanceOf(('fake identity') , { + from: '' + }, (err, output, data) => { + // console.log("contract call error: ", err) + // console.log('contract call data: ', data) + console.log('contract call output: ', output) + myContract.balanceOf(('fake identity') , { + from: '' + }, (err, output, data) => { + // console.log("contract call error: ", err) + // console.log('contract call data: ', data) + console.log('contract call output: ', output) + }); + }); + }); + }); + }); +}) diff --git a/minic_unittests/contracts/token_ant/token_ant.ds b/minic_unittests/contracts/token_ant/token_ant.ds new file mode 100644 index 0000000..e75015b --- /dev/null +++ b/minic_unittests/contracts/token_ant/token_ant.ds @@ -0,0 +1,47 @@ +(* Loosely based on the "sample fixed supply token contract" from https://theethereum.wiki/w/index.php/ERC20_Token_Standard#Sample_Fixed_Supply_Token_Contract *) + +external [[except DataTypeOps]] with "Require token.Invariant.\n" + +const _totalSupply = 100000 + +object signature ERC20Interface = { + constructor : unit -> unit; + const totalSupply : unit -> int; + const balanceOf : identity -> int; + transfer : identity * int -> bool +} + +object FixedSupplyToken : ERC20Interface { + let balances : mapping[identity] int := mapping_init + + let constructor () = + balances[msg_sender] := 100000 + + let totalSupply () = + let bal0 = balances[identity(0)] in + _totalSupply - bal0 + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + bal + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + if (fromA <> toA /\ from_bal >= tokens) + then begin + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + true + end else + false +} + +layer signature FIXEDSUPPLYTOKENSig = { + fixedSupplyToken : ERC20Interface +} + +layer FIXEDSUPPLYTOKEN : [{}] FIXEDSUPPLYTOKENSig = { + fixedSupplyToken = FixedSupplyToken +} assert "Invariant.inv" \ No newline at end of file diff --git a/minic_unittests/contracts/wasmtest/arith/arith.ds b/minic_unittests/contracts/wasmtest/arith/arith.ds new file mode 100644 index 0000000..467ee7c --- /dev/null +++ b/minic_unittests/contracts/wasmtest/arith/arith.ds @@ -0,0 +1,44 @@ +object signature ArithInterface = { + add : uint * uint -> uint; + sub : uint * uint -> uint; + mul : uint * uint -> uint; + div : uint * uint -> uint; + (* mod : uint * uint -> uint; *) +} + +object Arith : ArithInterface { + let add (n, m) = + let c = n + m in + assert (c >= n); + n + m + + let sub (n, m) = + assert (m <= n); + n - m + + let mul (n, m) = + if n = 0u0 then + 0u0 + else + let c = n * m in + let d = c / n in + assert (d = m); + c + + let div (n, m) = + assert (m > 0u0); + let c = n / m in + c + + (* let mod (n, m) = + assert (m != 0); + n mod m *) +} + +layer signature ARITHSig = { + arith : ArithInterface +} + +layer ARITH : [{}] ARITHSig = { + arith = Arith +} \ No newline at end of file diff --git a/minic_unittests/contracts/wasmtest/arith/arith.js b/minic_unittests/contracts/wasmtest/arith/arith.js new file mode 100755 index 0000000..a1877da --- /dev/null +++ b/minic_unittests/contracts/wasmtest/arith/arith.js @@ -0,0 +1,45 @@ +#!/usr/bin/env node +'strict'; +let main = require('../index'); + +if (process.argv.length != 6) { + console.log(`usage: ${process.argv[1]} ewasm-file func arg1 arg2`); + process.exit(0); +} + +function generate(funcSig, arg1, arg2) { + let callData = '' + switch (funcSig) { + case 'add': + callData += '771602f7'; + break; + case 'sub': + callData += 'b67d77c5'; + break; + case 'mul': + callData += 'c8a4ac9c'; + break; + case 'div': + callData += 'a391c15b'; + break; + // case 'mod': + // callData += 'f43f523a'; + // break; + // case 'shl': + // callData += '9da760ef'; + // break; + // case 'shr': + // callData += '75f4479a'; + // break; + default: + console.log(`unknown func: ${funcSig}`); + process.exit(0); + } + arg1 = BigInt(arg1); + callData += arg1.toString(16).padStart(64, '0'); + arg2 = BigInt(arg2); + callData += arg2.toString(16).padStart(64, '0'); + return callData; +} +let callData = generate(...process.argv.slice(3, 6)); +main(process.argv[2], callData).then(console.log).catch(console.log); diff --git a/unittests/asserts.ds b/minic_unittests/contracts/wasmtest/asserts/asserts.ds similarity index 100% rename from unittests/asserts.ds rename to minic_unittests/contracts/wasmtest/asserts/asserts.ds diff --git a/minic_unittests/contracts/wasmtest/asserts/asserts.js b/minic_unittests/contracts/wasmtest/asserts/asserts.js new file mode 100755 index 0000000..7b6901c --- /dev/null +++ b/minic_unittests/contracts/wasmtest/asserts/asserts.js @@ -0,0 +1,48 @@ +#!/usr/bin/env node +'strict'; +let main = require('../index'); + +if (process.argv.length != 6) { + console.log(`usage: ${process.argv[1]} ewasm-file func arg1 arg2`); + process.exit(0); +} + +function generate(funcSig, arg1, arg2) { + let callData = '' + switch (funcSig) { + case 'f': + callData += '771602f7'; + break; + case 'g': + callData += 'b67d77c5'; + break; + case 'h': + callData += 'c8a4ac9c'; + break; + case 'i': + callData += 'a391c15b'; + break; + case 'j': + callData += 'a391c15b'; + break; + // case 'mod': + // callData += 'f43f523a'; + // break; + // case 'shl': + // callData += '9da760ef'; + // break; + // case 'shr': + // callData += '75f4479a'; + // break; + default: + console.log(`unknown func: ${funcSig}`); + process.exit(0); + } + arg1 = BigInt(arg1); + callData += arg1.toString(16).padStart(64, '0'); + arg2 = BigInt(arg2); + callData += arg2.toString(16).padStart(64, '0'); + return callData; +} +let callData = generate(...process.argv.slice(3, 6)); +main(process.argv[2], callData).then(console.log).catch(console.log); diff --git a/minic_unittests/contracts/wasmtest/builtins/builtins.ds b/minic_unittests/contracts/wasmtest/builtins/builtins.ds new file mode 100644 index 0000000..e751e34 --- /dev/null +++ b/minic_unittests/contracts/wasmtest/builtins/builtins.ds @@ -0,0 +1,52 @@ +object signature OS = { + f : unit -> bool; + + const get_address : unit -> address; + const get_origin : unit -> address; + const get_caller : unit -> address; + const get_callvalue : unit -> uint; + const get_coinbase : unit -> uint; + const get_timestamp : unit -> uint; + const get_number : unit -> uint; + const get_balance_this : unit -> uint; + const get_blockhash_prev : unit -> uint +} + +object O : OS { + let _address : address := address(0) + let _origin : address := address(0) + let _caller : address := address(0) + let _callvalue : uint := 0u0 + let _coinbase : uint := 0u0 + let _timestamp : uint := 0u0 + let _number : uint := 0u0 + let _balance_this : uint := 0u0 + let _blockhash_prev : uint := 0u0 + + + let get_address () = let x = _address in x + let get_origin () = let x = _origin in x + let get_caller () = let x = _caller in x + let get_callvalue () = let x = _callvalue in x + let get_coinbase () = let x = _coinbase in x + let get_timestamp () = let x = _timestamp in x + let get_number () = let x = _number in x + let get_balance_this () = let x = _balance_this in x + let get_blockhash_prev () = let x = _blockhash_prev in x + + let f () = + _address := this_address; + _origin := tx_origin; + _caller := msg_sender; + _callvalue := msg_value; + _coinbase := block_coinbase; + _timestamp := block_timestamp; + _number := block_number; + _balance_this := balance(this_address); + _blockhash_prev := blockhash(block_number-0u1); + true + } + +layer L = { + o = O +} diff --git a/minic_unittests/contracts/wasmtest/builtins/builtins.js b/minic_unittests/contracts/wasmtest/builtins/builtins.js new file mode 100755 index 0000000..118927b --- /dev/null +++ b/minic_unittests/contracts/wasmtest/builtins/builtins.js @@ -0,0 +1,36 @@ +#!/usr/bin/env node +'strict'; +let main = require('../index'); + +if (process.argv.length != 4) { + console.log(`usage: ${process.argv[1]} ewasm-file func`); + process.exit(0); +} + +function generate(funcSig, arg1, arg2) { + let callData = '' + switch (funcSig) { + case 'f': + callData += '26121ff0'; + break; + // case 'mod': + // callData += 'f43f523a'; + // break; + // case 'shl': + // callData += '9da760ef'; + // break; + // case 'shr': + // callData += '75f4479a'; + // break; + default: + console.log(`unknown func: ${funcSig}`); + process.exit(0); + } + // arg1 = BigInt(arg1); + // callData += arg1.toString(16).padStart(64, '0'); + // arg2 = BigInt(arg2); + // callData += arg2.toString(16).padStart(64, '0'); + return callData; +} +let callData = generate(...process.argv.slice(3, 4)); +main(process.argv[2], callData).then(console.log).catch(console.log); diff --git a/minic_unittests/contracts/wasmtest/for/for.ds b/minic_unittests/contracts/wasmtest/for/for.ds new file mode 100644 index 0000000..9b4ad6e --- /dev/null +++ b/minic_unittests/contracts/wasmtest/for/for.ds @@ -0,0 +1,30 @@ +object signature OS = { + setThenGet : int -> int; + multiply : int * int -> int; +} + +object O : OS { + + let _val : int := 0 + + let setThenGet a = + _val := a; + let val = _val in + val + + (* a can be negative; + b must be positive *) + let multiply (a, b) = + for i = 0 to b do + begin + let val = _val in + _val := val + a + end; + let val = _val in + val + +} + +layer L = { + o = O +} diff --git a/minic_unittests/contracts/wasmtest/for/for.js b/minic_unittests/contracts/wasmtest/for/for.js new file mode 100755 index 0000000..49dcfea --- /dev/null +++ b/minic_unittests/contracts/wasmtest/for/for.js @@ -0,0 +1,29 @@ +#!/usr/bin/env node +'strict'; +let main = require('../index'); + +if (process.argv.length < 4) { + console.log(`usage: ${process.argv[1]} ewasm-file func [args..]`); + process.exit(0); +} + +function generate(funcSig, arg1, arg2) { + let callData = '' + switch (funcSig) { + case 'multiply': + callData += '165c4a16'; + break; + default: + console.log(`unknown func: ${funcSig}`); + process.exit(0); + } + arg1 = BigInt(arg1); + callData += arg1.toString(16).padStart(64, '0'); + if (arg2 !== undefined) { + arg2 = BigInt(arg2); + callData += arg2.toString(16).padStart(64, '0'); + } + return callData; +} +let callData = generate(...process.argv.slice(3, 6)); +main(process.argv[2], callData).then(console.log).catch(console.log); diff --git a/minic_unittests/contracts/wasmtest/index.js b/minic_unittests/contracts/wasmtest/index.js new file mode 100755 index 0000000..d88e551 --- /dev/null +++ b/minic_unittests/contracts/wasmtest/index.js @@ -0,0 +1,450 @@ +#!/usr/bin/env node + +'strict'; +let fs = require('fs'); +let utils = require('./utils'); +let precompiled = { + keccak256: null +}; +// const ethers = require("ethers"); + +class Environment { + constructor() { + this.storage = {}; + this.transactionReceipt = {}; + this.gasLeft = 65536; + this.callData = new Uint8Array(0); + this.returnData = new Uint8Array(0); + this.caller = '1234567890123456789012345678901234567890'; + this.callValue = '00000000000000000000000000000000'; + this.txGasPrice = '1f3f33ff5fabc5fdff67890feff12345'; + this.txOrigin = '1234567890123456789012345678901234567890'; + this.blockCoinbase = '1234567890123456789012345678901234567890'; + this.blockDifficulty = '0000000000000000000000000000000000000000000000000000405BBD86CA28'; + this.blockGasLimit = 21000; + this.blockNumber = 3456; + this.blockTimestamp = 6666; + this.blockhash = '1234567890123456789012345678901234567890123456789012345678901234'; + this.address_balance = '12345678901234567890123456789012'; + this.this_address = '5E72914535f202659083Db3a02C984188Fa26e9f'; + } + setCallData(callData) { + if (!/(^[0-9a-f]+)?/i.test(callData)) { + console.log(`except hex encoded calldata, got: ${callData}`); + return false; + } + if (callData.length > 0) { + if (callData.length % 2) + callData = "0" + callData; + this.callData = new Uint8Array(callData.length / 2).fill(0x00);; + callData.match(/.{2}/g).forEach((value, i) => { + this.callData[i] = parseInt(value, 16); + }); + } + return true; + } + getStorage(storage) { + return JSON.stringify(this.storage); + } + setStorage(storage) { + this.storage = JSON.parse(storage); + return true; + } +} + +class Interface { + get exports() { + let ret = { + ethereum: {}, + debug: {}, + }; + + // binding EEI functions + this.eei.forEach((method) => { + ret.ethereum[method] = this[method].bind(this); + }); + + // javascript doesn't support EEI with i64 parameter/return value, so binding to wrapper wasm function + this.hooks.forEach((method) => { + ret.ethereum[method] = this.eei_wrapper.instance.exports[method]; + }); + + [ + 'print32', + ].forEach((method) => { + ret.debug[method] = this[method].bind(this); + }); + + return ret; + } + constructor(env) { + this.env = env; + this.eei = [ + 'useGas', + 'getCallDataSize', + 'callDataCopy', + 'storageStore', + 'storageLoad', + 'log', + 'finish', + 'revert', + 'returnDataCopy', + 'getCaller', + 'getCallValue', + 'getTxGasPrice', + 'getTxOrigin', + 'getBlockCoinbase', + 'getBlockDifficulty', + 'returnDataSize', + 'getExternalBalance', + 'getAddress', + ]; + this.hooks = [ + 'getGasLeft', + 'callStatic', + 'call', + 'callDelegate', + 'getBlockGasLimit', + 'getBlockNumber', + 'getBlockHash', + 'getBlockTimestamp', + ]; + } + async connect() { + let hooks = {}; + this.hooks.forEach((method) => { + hooks[method] = this[method].bind(this); + }); + // EEI hook function injection + this.eei_wrapper = await WebAssembly.instantiate(fs.readFileSync("lib/wrapper.wasm"), { + ethereum: hooks + }); + } + getMemory(offset, length) { + return new Uint8Array(this.mem.buffer, offset, length); + } + setMemory(offset, length, value) { + const memory = new Uint8Array(this.mem.buffer, offset, length); + memory.set(value); + } + takeGas(amount) { + if (this.env.gasLeft < amount) { + throw new Error('Ran out of gas') + } + this.env.gasLeft -= amount + } + useGas(gas) { + console.log(`useGas(${gas})`); + //takeGas(gas); + } + getCallDataSize() { + console.log(`getCallDataSize()`); + this.takeGas(2); + console.log(`{ size: ${this.env.callData.length} }`); + return this.env.callData.length; + } + callDataCopy(resultOffset, dataOffset, length) { + console.log(`callDataCopy(${resultOffset}, ${dataOffset}, ${length})`); + this.takeGas(3 + Math.ceil(length / 32) * 3) + if (length) { + const callData = this.env.callData.slice(dataOffset, dataOffset + length); + this.setMemory(resultOffset, length, callData); + console.log(`{ data: ${utils.toHex(callData)} }`); + } + } + storageStore(pathOffset, valueOffset) { + console.log(`storageStore(${pathOffset}, ${valueOffset})`); + const path = utils.toBigInt('0x' + utils.toHex(this.getMemory(pathOffset, 32))).toString(16); + const value = utils.toBigInt('0x' + utils.toHex(this.getMemory(valueOffset, 32))).toString(16); + this.env.storage[path] = value; + console.log(`{ key: ${path}, value: ${value} }`); + } + storageLoad(pathOffset, valueOffset) { + console.log(`storageLoad(${pathOffset}, ${valueOffset})`); + const path = utils.toBigInt('0x' + utils.toHex(this.getMemory(pathOffset, 32))).toString(16); + if (path in this.env.storage) { + let value = this.env.storage[path]; + const data = value.padStart(64, '0').match(/.{2}/g).map(value => parseInt(value, 16)); + this.setMemory(valueOffset, 32, data); + console.log(`{ key: ${path}, value: ${value} }`); + } else { + const data = Array(32).fill(0); + this.setMemory(valueOffset, 32, data); + console.log(`{ key: ${path}, value: 0 }`); + } + } + log(dataOffset, dataLength, numberOfTopics, topic1, topic2, topic3, topic4) { + console.log(`log(${dataOffset}, ${dataLength}, ${numberOfTopics}, ${topic1}, ${topic2}, ${topic3}, ${topic4})`); + this.takeGas(375 + (375 * numberOfTopics) + (8 * dataLength)); + if (dataLength >= 1) { + const data = utils.toHex(this.getMemory(dataOffset, dataLength)); + this.env.transactionReceipt.data = data; + console.log(`{ data: ${data} }`); + } + if (numberOfTopics >= 1) { + this.env.transactionReceipt.topics = []; + const t1 = utils.toHex(this.getMemory(topic1, 32)); + this.env.transactionReceipt.topics.push(t1); + console.log(`{ t1/signature: ${t1} }`); + } + if (numberOfTopics >= 2) { + const t2 = utils.toHex(this.getMemory(topic2, 32)); + this.env.transactionReceipt.topics.push(t2); + console.log(`{ t2: ${t2} }`); + } + if (numberOfTopics >= 3) { + const t3 = utils.toHex(this.getMemory(topic3, 32)); + this.env.transactionReceipt.topics.push(t3); + console.log(`{ t3: ${t3} }`); + } + if (numberOfTopics >= 4) { + const t4 = utils.toHex(this.getMemory(topic4, 32)); + this.env.transactionReceipt.topics.push(t4); + console.log(`{ t4: ${t4} }`); + } + } + finish(dataOffset, dataLength) { + console.log(`finish(${dataOffset}, ${dataLength})`); + const data = this.getMemory(dataOffset, dataLength) + this.env.returnData = data; + throw new Error('finish'); + } + revert(dataOffset, dataLength) { + console.log(`revert(${dataOffset}, ${dataLength})`); + const data = this.getMemory(dataOffset, dataLength) + this.env.returnData = data; + const message = String.fromCharCode.apply(null, data); + this.env.transactionReceipt.revertReason = message; + console.log(`{ message: ${message} }`); + throw new Error('revert'); + process.exit(0); + } + callStatic(gas, addressOffset, dataOffset, dataLength) { + console.log(`callStatic(${gas}, ${addressOffset}, ${dataOffset}, ${dataLength})`); + const addressInHex = '0x' + utils.toHex(this.getMemory(addressOffset, 20)); + const address = utils.toBigInt(addressInHex); + const data = this.getMemory(dataOffset, dataLength); + console.log(`{ address: ${addressInHex}, data: ${utils.toHex(data)} }`); + + let vm; + switch (address) { + case BigInt(2): // Sha256 + case BigInt(9): // Keccak256 + vm = precompiled.keccak256; + break; + default: + return 1; + } + vm.run(utils.toHex(data)); + + this.env.returnData = vm.env.returnData; + return 0; + } + call(gas, addressOffset, valueOffset, dataOffset, dataLength) { + console.log(`call(${gas}, ${addressOffset}, ${valueOffset}, ${dataOffset}, ${dataLength})`); + const addressInHex = '0x' + utils.toHex(this.getMemory(addressOffset, 20)); + const address = utils.toBigInt(addressInHex); + const value = utils.toBigInt('0x' + utils.toHex(this.getMemory(valueOffset, 16))).toString(16); + const data = this.getMemory(dataOffset, dataLength); + console.log(`{ address: ${addressInHex}, value: ${value}, data: ${utils.toHex(data)} }`); + + let vm; + switch (address) { + case BigInt(2): // Sha256 + case BigInt(9): // Keccak256 + vm = precompiled.keccak256; + // ethers.utils.keccak256() + break; + default: + return 1; + } + vm.run(utils.toHex(data)); + + this.env.returnData = vm.env.returnData; + return 0; + } + callDelegate(gas, addressOffset, dataOffset, dataLength) { + console.log(`callDelegate(${gas}, ${addressOffset}, ${dataOffset}, ${dataLength})`); + const addressInHex = '0x' + utils.toHex(this.getMemory(addressOffset, 20)); + const address = utils.toBigInt(addressInHex); + const data = this.getMemory(dataOffset, dataLength); + console.log(`{ address: ${addressInHex}, data: ${utils.toHex(data)} }`); + + let vm; + switch (address) { + case BigInt(2): // Sha256 + case BigInt(9): // Keccak256 + vm = precompiled.keccak256; + break; + default: + return 1; + } + vm.run(utils.toHex(data)); + + this.env.returnData = vm.env.returnData; + return 0; + } + returnDataSize() { + console.log(`returnDataSize()`); + this.takeGas(2); + console.log(`{ size: ${this.env.returnData.length} }`); + return this.env.returnData.length; + } + returnDataCopy(resultOffset, dataOffset, length) { + console.log(`returnDataCopy(${resultOffset}, ${dataOffset}, ${length})`); + if (length) { + const callData = this.env.returnData.slice(dataOffset, dataOffset + length); + this.setMemory(resultOffset, length, callData); + console.log(`{ data: ${utils.toHex(callData)} }`) + } + } + getCaller(resultOffset) { + console.log(`getCaller(${resultOffset})`); + const data = this.env.caller.padStart(40, '0').match(/.{2}/g).map(value => parseInt(value, 16)); + this.setMemory(resultOffset, 20, data); + console.log(`{ caller: ${utils.toHex(data)} }`); + } + getCallValue(resultOffset) { + console.log(`getCallValue(${resultOffset})`); + const data = this.env.callValue.padStart(32, '0').match(/.{2}/g).map(value => parseInt(value, 16)); + this.setMemory(resultOffset, 16, data); + console.log(`{ value: ${utils.toHex(data)} }`); + } + getGasLeft() { + console.log(`getGasLeft()`); + console.log(`{ gas: ${this.env.gasLeft} }`); + return this.env.gasLeft; + } + + getTxGasPrice(valueOffset) { + console.log(`getTxGasPrice(${valueOffset})`); + const data = this.env.txGasPrice.padStart(32, '0').match(/.{2}/g).reverse().map(value => parseInt(value, 16)); + this.setMemory(valueOffset, 16, data); + console.log(`{ price: ${utils.toHex(data)} }`); + } + + getTxOrigin(resultOffset) { + console.log(`getTxOrigin(${resultOffset})`); + const data = this.env.txOrigin.padStart(40, '0').match(/.{2}/g).map(value => parseInt(value, 16)); + this.setMemory(resultOffset, 20, data); + console.log(`{ orig: ${utils.toHex(data)} }`); + } + + getBlockCoinbase(resultOffset) { + console.log(`getBlockCoinbase(${resultOffset})`); + const data = this.env.blockCoinbase.padStart(40, '0').match(/.{2}/g).map(value => parseInt(value, 16)); + this.setMemory(resultOffset, 20, data); + console.log(`{ coinbase: ${utils.toHex(data)} }`); + } + + getBlockDifficulty(resultOffset) { + console.log(`getBlockDifficulty(${resultOffset})`); + const data = this.env.blockDifficulty.padStart(64, '0').match(/.{2}/g).reverse().map(value => parseInt(value, 16)); + this.setMemory(resultOffset, 32, data); + console.log(`{ difficulty: ${utils.toHex(data)} }`); + } + + getBlockGasLimit() { + console.log(`getBlockGasLimit()`); + console.log(`{ gas: ${this.env.blockGasLimit} }`); + return this.env.blockGasLimit; + } + + getBlockNumber() { + console.log(`getBlockNumber()`); + console.log(`{ block: ${this.env.blockNumber} }`); + return this.env.blockNumber; + } + + getBlockTimestamp() { + console.log(`getBlockTimestamp()`); + console.log(`{ timestamp: ${this.env.blockTimestamp} }`); + return this.env.blockTimestamp; + } + + getBlockHash(number, resultOffset) { + console.log(`getBlockHash(${resultOffset})`); + const data = this.env.blockhash.padStart(64, '0').match(/.{2}/g).map(value => parseInt(value, 16)); + this.setMemory(resultOffset, 32, data); + console.log(`{ blockhash: ${utils.toHex(data)} }`); + return 0; + } + + getExternalBalance(addressOffset, resultOffset) { + console.log(`getExternalBalance(${addressOffset}, ${resultOffset})`); + const addressInHex = '0x' + utils.toHex(this.getMemory(addressOffset, 20)); + const address = utils.toBigInt(addressInHex); + const data = this.env.address_balance.padStart(32, '0').match(/.{2}/g).reverse().map(value => parseInt(value, 16)); + this.setMemory(resultOffset, 16, data); + console.log(`{ address: ${addressInHex}, value: ${utils.toHex(data)} }`); + } + + getAddress(resultOffset) { + console.log(`getAddress(${resultOffset})`); + const data = this.env.this_address.padStart(40, '0').match(/.{2}/g).map(value => parseInt(value, 16)); + this.setMemory(resultOffset, 20, data); + console.log(`{ orig: ${utils.toHex(data)} }`); + } + + print32(value) { + console.log(`print32(${(new Uint32Array([value]))[0]})`); + } +} + +class VM { + constructor(path) { + this.env = new Environment(); + this.int = new Interface(this.env); + this.path = path; + } + async instantiate() { + await this.int.connect(); + this.vm = await WebAssembly.instantiate(fs.readFileSync(this.path), this.int.exports); + this.int.mem = this.vm.instance.exports.memory; + } + run(callData, storage, env) { + if (!this.env.setCallData(callData)) { + return false; + } + if (storage && !this.env.setStorage(storage)) { + return false; + } + if (env) { + this.env = Object.assign(this.env, env); + } + try { + this.vm.instance.exports.main(); + } catch (error) { + if (error.message !== "finish") { + console.log(error); + } + } + return true; + } +} + +async function main(path, callData, storage, env) { + for (let name in precompiled) { + let vm = new VM(`lib/${name}.wasm`); + await vm.instantiate(); + precompiled[name] = vm; + } + let vm = new VM(path); + await vm.instantiate(); + vm.run(callData, storage, env); + return { + returnData: utils.toLEHex(vm.env.returnData), + storage: vm.env.getStorage(), + transactionReceipt: vm.env.transactionReceipt + }; +}; + +if (require.main === module) { + if (4 > process.argv.length || process.argv.length > 5) { + console.log(`usage: ${process.argv[1]} ewasm-file calldata [storage]`); + process.exit(0); + } + main(process.argv[2], process.argv[3], process.argv[4]).then(console.log).catch(console.log); +} else { + module.exports = main; +} \ No newline at end of file diff --git a/minic_unittests/contracts/wasmtest/utils.js b/minic_unittests/contracts/wasmtest/utils.js new file mode 100644 index 0000000..d23a150 --- /dev/null +++ b/minic_unittests/contracts/wasmtest/utils.js @@ -0,0 +1,14 @@ +module.exports = { + to256CallData: data => { + return BigInt(data).toString(16).padStart(64, '0'); + }, + toHex: data => { + return Array.from(data).map((value) => value.toString(16).padStart(2, '0')).join(''); + }, + toLEHex: data => { + return Array.from(data).map((value) => value.toString(16).padStart(2, '0')).reverse().join(''); + }, + toBigInt: hexRepr => { + return BigInt(hexRepr); + }, +} diff --git a/minic_unittests/declarations.mc b/minic_unittests/declarations.mc new file mode 100644 index 0000000..ed07ee3 --- /dev/null +++ b/minic_unittests/declarations.mc @@ -0,0 +1,21 @@ +int8 foo1; +int16 foo2; +int32 foo3; +int256 foo4; +bool foo5; +void foo6; + +// int256* foo7; +int256[256] foo8; +int256#[int256] foo9; +int256#[int256] foo10; + +struct bar1 { + int256 baz1; + bool baz2; +} foo11; + +union bar2 { + int256 baz1; + bool baz2; +} foo12; diff --git a/minic_unittests/expressions.mc b/minic_unittests/expressions.mc new file mode 100644 index 0000000..4c7fb63 --- /dev/null +++ b/minic_unittests/expressions.mc @@ -0,0 +1,42 @@ +int256[3] array; +int256#[bool] hashmap; +struct bar { + bool one; + int256#[int8] two; +} mystruct; + +void foo(int256 tmp) { + tmp = tmp; + tmp = 0; + // Pointer deref not supported by backend yet + // tmp = *tmp; + + tmp = !tmp; + tmp = ~tmp; + tmp = -tmp; + + tmp = 1 + 2 - 3 / 4 % 5 ** 6 & 7 | 8 ^ 9 << 10 >> 11 == 12; + + tmp = sha1 tmp; + + tmp = array[0]; + array[0] = tmp; + + tmp = hashmap[0]; + hashmap[0] = tmp; + + tmp = address; + tmp = origin; + tmp = caller; + tmp = callvalue; + tmp = coinbase; + tmp = timestamp; + tmp = number; + + tmp = balance(0); + tmp = blockhash(0); + + tmp = mystruct.one; + tmp = mystruct.two[0]; + +} diff --git a/minic_unittests/functions.mc b/minic_unittests/functions.mc new file mode 100644 index 0000000..1914de4 --- /dev/null +++ b/minic_unittests/functions.mc @@ -0,0 +1,5 @@ +private int8 foo1() {} +private int16 foo2() {} +private int32 foo3() {} +private int256 foo4() {} +private bool foo5() {} diff --git a/minic_unittests/methods.mc b/minic_unittests/methods.mc new file mode 100644 index 0000000..54da59b --- /dev/null +++ b/minic_unittests/methods.mc @@ -0,0 +1,5 @@ +int8 foo1() {} +int16 foo2() {} +int32 foo3() {} +int256 foo4() {} +bool foo5() {} diff --git a/minic_unittests/statements.mc b/minic_unittests/statements.mc new file mode 100644 index 0000000..7736e32 --- /dev/null +++ b/minic_unittests/statements.mc @@ -0,0 +1,60 @@ +private int256 foo() { + int256 tmp; + + tmp = 0; + return tmp; +} + +void bar(int256 tmp) { + tmp = 0; + transfer(0, 1); + + foo(); + tmp <- foo(); + + { + tmp = 0; + } + + if (tmp == 0) { + } + + if (tmp == 0) + tmp = 1; + + if (tmp == 0) + tmp = 1; + else + tmp = 2; + + if (tmp == 0) { + tmp = 1; + } else { + tmp = 2; + } + + while + tmp = tmp - 1; + + while { + tmp = tmp - 1; + } + + while { + tmp = tmp - 1; + foo(); + if (tmp == 0) break; + } + + while { + tmp <- foo(); + if (tmp == 0) break; + } + + callmethod(tmp; tmp; 12345; 6; ; 7); + + revert; + return; + return tmp; + +} diff --git a/minic_unittests/test_bytecode.sh b/minic_unittests/test_bytecode.sh new file mode 100755 index 0000000..568341e --- /dev/null +++ b/minic_unittests/test_bytecode.sh @@ -0,0 +1,44 @@ +#!/bin/bash + +fail_count=0 +pass_count=0 + +minicc=../binaries/MacOS/minicc +dsc=../binaries/MacOS/dsc + +skip_files="token_ant.ds" + +edsger_out_mc=$(mktemp) +edsger_out_asmbl=$(mktemp) +minicc_out_mc=$(mktemp) +minicc_out_asmbl=$(mktemp) + +on_exit() { + echo -e '\nCleaning up...' + rm $edsger_out_mc $edsger_out_asmbl \ + $minicc_out_mc $minicc_out_asmbl + echo $pass_count passed, $fail_count failed +} +trap on_exit EXIT + +for f in ./unittests/*.ds ./contracts/*/*.ds; do + bname=$(basename $f) + if echo $skip_files | grep -q -w $bname; then + continue + fi + + echo -n $f + + $dsc $f minic > $edsger_out_mc + $dsc $f assembly > $edsger_out_asmbl + $minicc $edsger_out_mc print > $minicc_out_mc + $minicc $edsger_out_mc assembly > $minicc_out_asmbl + + if cmp -s $edsger_out_asmbl $minicc_out_asmbl; then + echo + ((pass_count++)) + else + echo : FAIL + ((fail_count++)) + fi +done diff --git a/minic_unittests/test_parse_and_print.sh b/minic_unittests/test_parse_and_print.sh new file mode 100755 index 0000000..c78b708 --- /dev/null +++ b/minic_unittests/test_parse_and_print.sh @@ -0,0 +1,52 @@ +#!/bin/bash + +# Add deepsea (.ds) or minic (.mc) files to this directory to add new unit +# tests. Deepsea files are first prett-printed to minic source. Then, all minic +# files are parsed and printed twice and checked for any differences. + +minicc=../binaries/MacOS/minicc +dsc=../binaries/MacOS/dsc + +skip_files="token_ant.ds" + +deepsea_compiled=$(mktemp -d) +tmp1=$(mktemp) +tmp2=$(mktemp) + +on_exit () { + echo -e '\nCleaning up...' + rm -r "$deepsea_compiled" + rm $tmp1 $tmp2 + echo $pass_count passed, $fail_count failed +} +trap on_exit EXIT + +for f in ./unittests/*.ds ./contracts/*/*.ds; do + bname=$(basename $f) + if echo $skip_files | grep -q -w $bname; then + continue + fi + echo Compiling $bname + + $dsc $f minic > $deepsea_compiled/${bname%.*}.mc +done + +echo -e '\nRunning tests' + +fail_count=0 +pass_count=0 + +for mc in ./*.mc $deepsea_compiled/*; do + echo -n $(basename $mc) + + $minicc $mc > $tmp1 + $minicc $tmp1 > $tmp2 + + if diff -u $tmp1 $tmp2; then + echo + ((pass_count++)) + else + echo : FAIL + ((fail_count++)) + fi +done diff --git a/unittests/address.ds b/minic_unittests/unittests/address.ds similarity index 100% rename from unittests/address.ds rename to minic_unittests/unittests/address.ds diff --git a/minic_unittests/unittests/amm.ds b/minic_unittests/unittests/amm.ds new file mode 100644 index 0000000..b5bb763 --- /dev/null +++ b/minic_unittests/unittests/amm.ds @@ -0,0 +1,401 @@ +(* TODO: add emit Events *) + +(* PROOFS: + 1. change re-defined layers + 2. balance tracks reserve + 3. increasing k -> algebra + 4. no path independence, splitting trades is strictly more expensive -> algebra, appendix D + 5. liquidity token economics is right, i.e., UNI-V2 tracks the actual share of pool --> actually, where is the trading of UNI-V2 Pair tokens? What is its value? + 6. MOD: slippage control + 7. ... cost of manipulation (hard) +*) + +object signature ERC20Interface = { + constructor : unit -> unit; + mint : address * int -> unit; + burn : address * int -> unit; + totalSupply : unit -> int; + balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object LiquidityToken : ERC20Interface { + let _totalSupply : int := 0 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + balances[msg_sender] := 100000 + + let mint (toA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply + value; + let to_bal = balances[toA] in + balances[toA] := to_bal + value + + let burn (fromA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply - value; + let from_bal = balances[fromA] in + balances[fromA] := from_bal - value + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +object signature FixedERC20Interface = { + constructor : unit -> unit; + const totalSupply : unit -> int; + const balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object FixedSupplyToken : FixedERC20Interface { + let _totalSupply : int := 100000 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + _totalSupply := 100000; + balances[msg_sender] := 100000 + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +(* this AMM supports ERC20 tokens, ETH can be incorporated as WETH for trading *) +object signature AMMInterface = { + constructor : unit -> unit; + (* swap : int * int * address -> unit; *) + + coarsed simpleSwap0 : address -> int; (* transfer token0 to amm, results number of token1 swapped *) + + (* router *) + (* swapTokensForExactTokens + swapExactTokensForTokens + addLiquidity (* why restrain the liquidity to not depreciating an asset too much? *) + removeLiquidity *) + + const getReserve0 : unit -> int; + const getReserve1 : unit -> int; + const getBalance0 : unit -> int; + + (* fund management functions *) + coarsed mint : address -> unit; (* provide liquidity to this pair *) + coarsed burn : address -> unit; (* remove liquidity from this pair *) + + coarsed skim : address -> unit; (* arbitrage to force balances to match reserves *) + coarsed sync : unit -> unit; (* force reserves to match balances *) + k : unit -> int; (* get the constant product *) + + (* oracles *) + (* we do not provide Uniswap V2 version of smoothed oracles at this moment *) + quote0 : int -> int; (* given some amount of an asset and pair reserves, returns an equivalent amount of the other asset *) + getAmountOut0 : int -> int; (* given an input amount of an asset and pair reserves, returns the maximum output amount of the other asset *) + getAmountIn0 : int -> int; (* given an output amount of an asset and pair reserves, returns a required input amount of the other asset *) +} + +(* assuming 0.3% txn fee, excluding the 0.05% additional fee in Uniswap V2 design *) +(* we do not support liquidity tokens at the moment *) +object AutomatedMarketMaker + ( liquidityToken: ERC20Interface, + erc20Token0: FixedERC20Interface, + erc20Token1: FixedERC20Interface) : AMMInterface { + + let _token0 : address := address(0xdac17f958d2ee523a2206206994597c13d831ec7) (* USDT *) + let _token1 : address := address(0x6b175474e89094c44da98b954eedeac495271d0f) (* DAI *) + let _owner : address := address(0x0) + let _reserve0 : int := 0 + let _reserve1 : int := 0 + let blockTimestampLast : int := 0 (* new oracle *) + let price0CumulativeLast : int := 0 (* new oracle *) + let price1CumulativeLast : int := 0 (* new oracle *) + let kLast : int := 0 (* reserve0 * reserve1, as of immediately after the most recent liquidity event *) + + let constructor () = + _owner := msg_sender + + let getReserve0 () = + let reserve0 = _reserve0 in + reserve0 + + let getReserve1 () = + let reserve1 = _reserve1 in + reserve1 + + let getBalance0 () = + let balance0 = erc20Token0.balanceOf(this_address) in + balance0 + + (* transfer token0 and token1 to this contract, then mint liquidity tokens *) + let mint (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0 = balance0 - reserve0 in + let amount1 = balance1 - reserve1 in + (* update reserve 0 and reserve 1 *) + let totalSupply = liquidityToken.totalSupply() in + let liquidity = if totalSupply = 0 + then + begin + liquidityToken.mint(address(0x0), 1000); (* lock the first liquidity tokens*) + amount0 * amount1 - 1000 (* in Uniswap this should be square-rooted *) + end + else + let x = amount0 * totalSupply / reserve0 in + let y = amount1 * totalSupply / reserve1 in + if x < y then x else y + in + assert (liquidity > 0); + liquidityToken.mint(toA, liquidity); + _reserve0 := balance0; + _reserve1 := balance1 + + (* need toA first transfer liquidity tokens to this contract, then operate *) + (* it does not use ERC20 style allowances, just direct transfer *) + let burn (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let liquidity = liquidityToken.balanceOf(this_address) in + let totalSupply = liquidityToken.totalSupply() in + let amount0 = liquidity * balance0 / totalSupply in + let amount1 = liquidity * balance1 / totalSupply in + assert (amount0 > 0 /\ amount1 > 0); + liquidityToken.burn(this_address, liquidity); + let success = erc20Token0.transfer(toA, amount0) in + assert (success); + let success = erc20Token1.transfer(toA, amount1) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + (* adhere to the stanford paper formalization, targeted for vanilla proof *) + let simpleSwap0 (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = balance0 - reserve0 in + let token0 = _token0 in + let token1 = _token1 in + assert (toA <> token0 /\ toA <> token1); + assert (amount0In > 0); + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let result = numerator / denominator in + let success = erc20Token1.transfer(toA, result) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1; + let resultU = result in + resultU + + (* derives from Uniswap V2 flashswap, but not flashswap since DeepSEA does not allow control-flow aggregation *) + (* first transfer tokens to this contract, then call swap to swap them *) + (* notice this does not refund, so possibly not optimal rate *) + (* let swap (amount0Out, amount1Out, toA) = + assert (amount0Out > 0 \/ amount1Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let token0 = _token0 in + let token1 = _token1 in + assert (amount0Out < reserve0 /\ amount1Out < reserve1); + assert (toA <> token0 /\ toA <> token1); + (* check the assets satisfy what is wanted *) + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + (* we can prove that balance > reserve here *) + let amount0In = helper.getAmountIn(balance0, 0, reserve0) in (* balance0 - reserve0 *) + let amount1In = helper.getAmountIn(balance1, 0, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0 - amount0Out, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1 - amount1Out, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); + + (* sending money *) + let _ = if amount0Out > 0 then + let success = erc20Token0.transfer(toA, amount0Out) in + assert (success); + () + else + () + in + let _ = if amount1Out > 0 then + let success = erc20Token1.transfer(toA, amount1Out) in + assert (success); + () + else + () + in + (* flash swap style: sent money (amount0Out, amount1Out) to msg_sender, now wait to get back *) + (* after got money back *) + (* let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = helper.getAmountIn(balance0, amount0Out, reserve0) in + let amount1In = helper.getAmountIn(balance1, amount1Out, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); *) + (* update the reserves, also update the oracle if possible *) + _reserve0 := balance0; + _reserve1 := balance1 *) + + let skim (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let skim0 = balance0 - reserve0 in + let skim1 = balance1 - reserve1 in + let success = erc20Token0.transfer(toA, skim0) in + assert (success); + let success = erc20Token1.transfer(toA, skim1) in + assert (success) + + let sync () = + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + let k () = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let resultU = reserve0 * reserve1 in + resultU + + (* Uniswap did not use Q112 library to do this computation *) + + (* given amount0 of token0, the equivalent value in token1 *) + let quote0 (amount0) = + assert (amount0 > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); (* prove that this can be dropped? *) + let resultU = amount0 * reserve1 / reserve0 in + resultU + + let getAmountOut0 (amount0In) = + assert (amount0In > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let resultU = numerator / denominator in + resultU + + let getAmountIn0 (amount0Out) = + assert (amount0Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let numerator = reserve1 * amount0Out * 1000 in + let denominator = (reserve0 - amount0Out) * 997 in + let resultU = (numerator / denominator) + 1 in + resultU +} + +layer signature AMMLibSig = { + erc20Token0 : FixedERC20Interface; + erc20Token1 : FixedERC20Interface; + liquidityToken : ERC20Interface; +} + +layer AMMLIB : [{}] AMMLibSig = { + erc20Token0 = address(0x10033) <: FixedSupplyToken; + erc20Token1 = address(0x10032) <: (clone FixedSupplyToken); (* cloned an object *) + liquidityToken = address(0x10031) <: LiquidityToken; +} + +layer signature AMMSig = { + amm : AMMInterface +} + +layer AMM : [AMMLibSig] AMMSig = { + amm = address(0x10030) <: AutomatedMarketMaker +} + +layer COMPLETE = AMM @ AMMLIB diff --git a/minic_unittests/unittests/amm_with_addresses.ds b/minic_unittests/unittests/amm_with_addresses.ds new file mode 100644 index 0000000..5c8ce0a --- /dev/null +++ b/minic_unittests/unittests/amm_with_addresses.ds @@ -0,0 +1,401 @@ +(* TODO: add emit Events *) + +(* PROOFS: + 1. change re-defined layers + 2. balance tracks reserve + 3. increasing k -> algebra + 4. no path independence, splitting trades is strictly more expensive -> algebra, appendix D + 5. liquidity token economics is right, i.e., UNI-V2 tracks the actual share of pool --> actually, where is the trading of UNI-V2 Pair tokens? What is its value? + 6. MOD: slippage control + 7. ... cost of manipulation (hard) +*) + +object signature ERC20Interface = { + constructor : unit -> unit; + mint : address * int -> unit; + burn : address * int -> unit; + totalSupply : unit -> int; + balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object LiquidityToken : ERC20Interface { + let _totalSupply : int := 0 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + balances[msg_sender] := 100000 + + let mint (toA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply + value; + let to_bal = balances[toA] in + balances[toA] := to_bal + value + + let burn (fromA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply - value; + let from_bal = balances[fromA] in + balances[fromA] := from_bal - value + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +object signature FixedERC20Interface = { + constructor : unit -> unit; + const totalSupply : unit -> int; + const balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object FixedSupplyToken : FixedERC20Interface { + let _totalSupply : int := 100000 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + _totalSupply := 100000; + balances[msg_sender] := 100000 + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +(* this AMM supports ERC20 tokens, ETH can be incorporated as WETH for trading *) +object signature AMMInterface = { + constructor : unit -> unit; + (* swap : int * int * address -> unit; *) + + coarsed simpleSwap0 : address -> int; (* transfer token0 to amm, results number of token1 swapped *) + + (* router *) + (* swapTokensForExactTokens + swapExactTokensForTokens + addLiquidity (* why restrain the liquidity to not depreciating an asset too much? *) + removeLiquidity *) + + const getReserve0 : unit -> int; + const getReserve1 : unit -> int; + const getBalance0 : unit -> int; + + (* fund management functions *) + coarsed mint : address -> unit; (* provide liquidity to this pair *) + coarsed burn : address -> unit; (* remove liquidity from this pair *) + + coarsed skim : address -> unit; (* arbitrage to force balances to match reserves *) + coarsed sync : unit -> unit; (* force reserves to match balances *) + k : unit -> int; (* get the constant product *) + + (* oracles *) + (* we do not provide Uniswap V2 version of smoothed oracles at this moment *) + quote0 : int -> int; (* given some amount of an asset and pair reserves, returns an equivalent amount of the other asset *) + getAmountOut0 : int -> int; (* given an input amount of an asset and pair reserves, returns the maximum output amount of the other asset *) + getAmountIn0 : int -> int; (* given an output amount of an asset and pair reserves, returns a required input amount of the other asset *) +} + +(* assuming 0.3% txn fee, excluding the 0.05% additional fee in Uniswap V2 design *) +(* we do not support liquidity tokens at the moment *) +object AutomatedMarketMaker + ( liquidityToken: ERC20Interface, + erc20Token0: FixedERC20Interface, + erc20Token1: FixedERC20Interface) : AMMInterface { + + let _token0 : address := address(0xdac17f958d2ee523a2206206994597c13d831ec7) (* USDT *) + let _token1 : address := address(0x6b175474e89094c44da98b954eedeac495271d0f) (* DAI *) + let _owner : address := address(0x0) + let _reserve0 : int := 0 + let _reserve1 : int := 0 + let blockTimestampLast : int := 0 (* new oracle *) + let price0CumulativeLast : int := 0 (* new oracle *) + let price1CumulativeLast : int := 0 (* new oracle *) + let kLast : int := 0 (* reserve0 * reserve1, as of immediately after the most recent liquidity event *) + + let constructor () = + _owner := msg_sender + + let getReserve0 () = + let reserve0 = _reserve0 in + reserve0 + + let getReserve1 () = + let reserve1 = _reserve1 in + reserve1 + + let getBalance0 () = + let balance0 = erc20Token0.balanceOf(this_address) in + balance0 + + (* transfer token0 and token1 to this contract, then mint liquidity tokens *) + let mint (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0 = balance0 - reserve0 in + let amount1 = balance1 - reserve1 in + (* update reserve 0 and reserve 1 *) + let totalSupply = liquidityToken.totalSupply() in + let liquidity = if totalSupply = 0 + then + begin + liquidityToken.mint(address(0x0), 1000); (* lock the first liquidity tokens*) + amount0 * amount1 - 1000 (* in Uniswap this should be square-rooted *) + end + else + let x = amount0 * totalSupply / reserve0 in + let y = amount1 * totalSupply / reserve1 in + if x < y then x else y + in + assert (liquidity > 0); + liquidityToken.mint(toA, liquidity); + _reserve0 := balance0; + _reserve1 := balance1 + + (* need toA first transfer liquidity tokens to this contract, then operate *) + (* it does not use ERC20 style allowances, just direct transfer *) + let burn (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let liquidity = liquidityToken.balanceOf(this_address) in + let totalSupply = liquidityToken.totalSupply() in + let amount0 = liquidity * balance0 / totalSupply in + let amount1 = liquidity * balance1 / totalSupply in + assert (amount0 > 0 /\ amount1 > 0); + liquidityToken.burn(this_address, liquidity); + let success = erc20Token0.transfer(toA, amount0) in + assert (success); + let success = erc20Token1.transfer(toA, amount1) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + (* adhere to the stanford paper formalization, targeted for vanilla proof *) + let simpleSwap0 (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = balance0 - reserve0 in + let token0 = _token0 in + let token1 = _token1 in + assert (toA <> token0 /\ toA <> token1); + assert (amount0In > 0); + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let result = numerator / denominator in + let success = erc20Token1.transfer(toA, result) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1; + let resultU = result in + resultU + + (* derives from Uniswap V2 flashswap, but not flashswap since DeepSEA does not allow control-flow aggregation *) + (* first transfer tokens to this contract, then call swap to swap them *) + (* notice this does not refund, so possibly not optimal rate *) + (* let swap (amount0Out, amount1Out, toA) = + assert (amount0Out > 0 \/ amount1Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let token0 = _token0 in + let token1 = _token1 in + assert (amount0Out < reserve0 /\ amount1Out < reserve1); + assert (toA <> token0 /\ toA <> token1); + (* check the assets satisfy what is wanted *) + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + (* we can prove that balance > reserve here *) + let amount0In = helper.getAmountIn(balance0, 0, reserve0) in (* balance0 - reserve0 *) + let amount1In = helper.getAmountIn(balance1, 0, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0 - amount0Out, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1 - amount1Out, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); + + (* sending money *) + let _ = if amount0Out > 0 then + let success = erc20Token0.transfer(toA, amount0Out) in + assert (success); + () + else + () + in + let _ = if amount1Out > 0 then + let success = erc20Token1.transfer(toA, amount1Out) in + assert (success); + () + else + () + in + (* flash swap style: sent money (amount0Out, amount1Out) to msg_sender, now wait to get back *) + (* after got money back *) + (* let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = helper.getAmountIn(balance0, amount0Out, reserve0) in + let amount1In = helper.getAmountIn(balance1, amount1Out, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); *) + (* update the reserves, also update the oracle if possible *) + _reserve0 := balance0; + _reserve1 := balance1 *) + + let skim (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let skim0 = balance0 - reserve0 in + let skim1 = balance1 - reserve1 in + let success = erc20Token0.transfer(toA, skim0) in + assert (success); + let success = erc20Token1.transfer(toA, skim1) in + assert (success) + + let sync () = + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + let k () = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let resultU = reserve0 * reserve1 in + resultU + + (* Uniswap did not use Q112 library to do this computation *) + + (* given amount0 of token0, the equivalent value in token1 *) + let quote0 (amount0) = + assert (amount0 > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); (* prove that this can be dropped? *) + let resultU = amount0 * reserve1 / reserve0 in + resultU + + let getAmountOut0 (amount0In) = + assert (amount0In > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let resultU = numerator / denominator in + resultU + + let getAmountIn0 (amount0Out) = + assert (amount0Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let numerator = reserve1 * amount0Out * 1000 in + let denominator = (reserve0 - amount0Out) * 997 in + let resultU = (numerator / denominator) + 1 in + resultU +} + +layer signature AMMLibSig = { + erc20Token0 : FixedERC20Interface; + erc20Token1 : FixedERC20Interface; + liquidityToken : ERC20Interface; +} + +layer AMMLIB : [{}] AMMLibSig = { + erc20Token0 = address(0xB41b77dceD744699bc1a49807e2C41280958F5a2) <: FixedSupplyToken; + erc20Token1 = address(0x941b185dE7183e6cc00207C67E241e15ceCd83eE) <: (clone FixedSupplyToken); (* cloned an object *) + liquidityToken = address(0x5868FeaBCdAe035a5A652162C98d5edE2D90df1c) <: LiquidityToken; +} + +layer signature AMMSig = { + amm : AMMInterface +} + +layer AMM : [AMMLibSig] AMMSig = { + amm = address(0x10030) <: AutomatedMarketMaker +} + +layer COMPLETE = AMM @ AMMLIB diff --git a/minic_unittests/unittests/asserts.ds b/minic_unittests/unittests/asserts.ds new file mode 100755 index 0000000..4b72626 --- /dev/null +++ b/minic_unittests/unittests/asserts.ds @@ -0,0 +1,34 @@ +object signature OS = { + f : unit -> unit; + g : unit -> unit; + h : unit -> unit; + i : unit -> unit; + j : unit -> unit +} + +object O : OS { + (* should succeed *) + let f () = + assert (2+2 = 4) + + (* should revert *) + let g () = + assert (2+2 = 5) + + (* should revert *) + let h () = + deny (2+2 = 4) + + (* should succeed *) + let i () = + deny (2+2 = 5) + + (* should revert *) + let j () = + fail + +} + +layer L = { + o = O +} diff --git a/unittests/builtins.ds b/minic_unittests/unittests/builtins.ds old mode 100644 new mode 100755 similarity index 100% rename from unittests/builtins.ds rename to minic_unittests/unittests/builtins.ds diff --git a/unittests/constructor.ds b/minic_unittests/unittests/constructor.ds similarity index 100% rename from unittests/constructor.ds rename to minic_unittests/unittests/constructor.ds diff --git a/unittests/constructor2.ds b/minic_unittests/unittests/constructor2.ds similarity index 100% rename from unittests/constructor2.ds rename to minic_unittests/unittests/constructor2.ds diff --git a/minic_unittests/unittests/erc20Token0.ds b/minic_unittests/unittests/erc20Token0.ds new file mode 100644 index 0000000..a7b8685 --- /dev/null +++ b/minic_unittests/unittests/erc20Token0.ds @@ -0,0 +1,404 @@ +(* TODO: add emit Events *) + +(* PROOFS: + 1. change re-defined layers + 2. balance tracks reserve + 3. increasing k -> algebra + 4. no path independence, splitting trades is strictly more expensive -> algebra, appendix D + 5. liquidity token economics is right, i.e., UNI-V2 tracks the actual share of pool --> actually, where is the trading of UNI-V2 Pair tokens? What is its value? + 6. MOD: slippage control + 7. ... cost of manipulation (hard) +*) + +object signature ERC20Interface = { + constructor : unit -> unit; + mint : address * int -> unit; + burn : address * int -> unit; + totalSupply : unit -> int; + balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object LiquidityToken : ERC20Interface { + let _totalSupply : int := 0 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + balances[msg_sender] := 100000 + + let mint (toA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply + value; + let to_bal = balances[toA] in + balances[toA] := to_bal + value + + let burn (fromA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply - value; + let from_bal = balances[fromA] in + balances[fromA] := from_bal - value + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +object signature FixedERC20Interface = { + constructor : unit -> unit; + const totalSupply : unit -> int; + const balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object FixedSupplyToken : FixedERC20Interface { + let _totalSupply : int := 100000 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + _totalSupply := 100000; + balances[msg_sender] := 100000 + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +(* this AMM supports ERC20 tokens, ETH can be incorporated as WETH for trading *) +object signature AMMInterface = { + constructor : unit -> unit; + (* swap : int * int * address -> unit; *) + + coarsed simpleSwap0 : address -> int; (* transfer token0 to amm, results number of token1 swapped *) + + (* router *) + (* swapTokensForExactTokens + swapExactTokensForTokens + addLiquidity (* why restrain the liquidity to not depreciating an asset too much? *) + removeLiquidity *) + + const getReserve0 : unit -> int; + const getReserve1 : unit -> int; + const getBalance0 : unit -> int; + + (* fund management functions *) + coarsed mint : address -> unit; (* provide liquidity to this pair *) + coarsed burn : address -> unit; (* remove liquidity from this pair *) + + coarsed skim : address -> unit; (* arbitrage to force balances to match reserves *) + coarsed sync : unit -> unit; (* force reserves to match balances *) + k : unit -> int; (* get the constant product *) + + (* oracles *) + (* we do not provide Uniswap V2 version of smoothed oracles at this moment *) + quote0 : int -> int; (* given some amount of an asset and pair reserves, returns an equivalent amount of the other asset *) + getAmountOut0 : int -> int; (* given an input amount of an asset and pair reserves, returns the maximum output amount of the other asset *) + getAmountIn0 : int -> int; (* given an output amount of an asset and pair reserves, returns a required input amount of the other asset *) +} + +(* assuming 0.3% txn fee, excluding the 0.05% additional fee in Uniswap V2 design *) +(* we do not support liquidity tokens at the moment *) +object AutomatedMarketMaker + ( liquidityToken: ERC20Interface, + erc20Token0: FixedERC20Interface, + erc20Token1: FixedERC20Interface) : AMMInterface { + + let _token0 : address := address(0xdac17f958d2ee523a2206206994597c13d831ec7) (* USDT *) + let _token1 : address := address(0x6b175474e89094c44da98b954eedeac495271d0f) (* DAI *) + let _owner : address := address(0x0) + let _reserve0 : int := 0 + let _reserve1 : int := 0 + let blockTimestampLast : int := 0 (* new oracle *) + let price0CumulativeLast : int := 0 (* new oracle *) + let price1CumulativeLast : int := 0 (* new oracle *) + let kLast : int := 0 (* reserve0 * reserve1, as of immediately after the most recent liquidity event *) + + let constructor () = + _owner := msg_sender + + let getReserve0 () = + let reserve0 = _reserve0 in + reserve0 + + let getReserve1 () = + let reserve1 = _reserve1 in + reserve1 + + let getBalance0 () = + let balance0 = erc20Token0.balanceOf(this_address) in + balance0 + + (* transfer token0 and token1 to this contract, then mint liquidity tokens *) + let mint (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0 = balance0 - reserve0 in + let amount1 = balance1 - reserve1 in + (* update reserve 0 and reserve 1 *) + let totalSupply = liquidityToken.totalSupply() in + let liquidity = if totalSupply = 0 + then + begin + liquidityToken.mint(address(0x0), 1000); (* lock the first liquidity tokens*) + amount0 * amount1 - 1000 (* in Uniswap this should be square-rooted *) + end + else + let x = amount0 * totalSupply / reserve0 in + let y = amount1 * totalSupply / reserve1 in + if x < y then x else y + in + assert (liquidity > 0); + liquidityToken.mint(toA, liquidity); + _reserve0 := balance0; + _reserve1 := balance1 + + (* need toA first transfer liquidity tokens to this contract, then operate *) + (* it does not use ERC20 style allowances, just direct transfer *) + let burn (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let liquidity = liquidityToken.balanceOf(this_address) in + let totalSupply = liquidityToken.totalSupply() in + let amount0 = liquidity * balance0 / totalSupply in + let amount1 = liquidity * balance1 / totalSupply in + assert (amount0 > 0 /\ amount1 > 0); + liquidityToken.burn(this_address, liquidity); + let success = erc20Token0.transfer(toA, amount0) in + assert (success); + let success = erc20Token1.transfer(toA, amount1) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + (* adhere to the stanford paper formalization, targeted for vanilla proof *) + let simpleSwap0 (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = balance0 - reserve0 in + let token0 = _token0 in + let token1 = _token1 in + assert (toA <> token0 /\ toA <> token1); + assert (amount0In > 0); + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let result = numerator / denominator in + let success = erc20Token1.transfer(toA, result) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1; + let resultU = result in + resultU + + (* derives from Uniswap V2 flashswap, but not flashswap since DeepSEA does not allow control-flow aggregation *) + (* first transfer tokens to this contract, then call swap to swap them *) + (* notice this does not refund, so possibly not optimal rate *) + (* let swap (amount0Out, amount1Out, toA) = + assert (amount0Out > 0 \/ amount1Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let token0 = _token0 in + let token1 = _token1 in + assert (amount0Out < reserve0 /\ amount1Out < reserve1); + assert (toA <> token0 /\ toA <> token1); + (* check the assets satisfy what is wanted *) + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + (* we can prove that balance > reserve here *) + let amount0In = helper.getAmountIn(balance0, 0, reserve0) in (* balance0 - reserve0 *) + let amount1In = helper.getAmountIn(balance1, 0, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0 - amount0Out, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1 - amount1Out, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); + + (* sending money *) + let _ = if amount0Out > 0 then + let success = erc20Token0.transfer(toA, amount0Out) in + assert (success); + () + else + () + in + let _ = if amount1Out > 0 then + let success = erc20Token1.transfer(toA, amount1Out) in + assert (success); + () + else + () + in + (* flash swap style: sent money (amount0Out, amount1Out) to msg_sender, now wait to get back *) + (* after got money back *) + (* let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = helper.getAmountIn(balance0, amount0Out, reserve0) in + let amount1In = helper.getAmountIn(balance1, amount1Out, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); *) + (* update the reserves, also update the oracle if possible *) + _reserve0 := balance0; + _reserve1 := balance1 *) + + let skim (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let skim0 = balance0 - reserve0 in + let skim1 = balance1 - reserve1 in + let success = erc20Token0.transfer(toA, skim0) in + assert (success); + let success = erc20Token1.transfer(toA, skim1) in + assert (success) + + let sync () = + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + let k () = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let resultU = reserve0 * reserve1 in + resultU + + (* Uniswap did not use Q112 library to do this computation *) + + (* given amount0 of token0, the equivalent value in token1 *) + let quote0 (amount0) = + assert (amount0 > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); (* prove that this can be dropped? *) + let resultU = amount0 * reserve1 / reserve0 in + resultU + + let getAmountOut0 (amount0In) = + assert (amount0In > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let resultU = numerator / denominator in + resultU + + let getAmountIn0 (amount0Out) = + assert (amount0Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let numerator = reserve1 * amount0Out * 1000 in + let denominator = (reserve0 - amount0Out) * 997 in + let resultU = (numerator / denominator) + 1 in + resultU +} + +layer signature AMMLibSig = { + erc20Token0 : FixedERC20Interface; + erc20Token1 : FixedERC20Interface; + liquidityToken : ERC20Interface; +} + +layer AMMLIB : [{}] AMMLibSig = { + erc20Token0 = address(0x10033) <: FixedSupplyToken; + erc20Token1 = address(0x10032) <: (clone FixedSupplyToken); (* cloned an object *) + liquidityToken = address(0x10031) <: LiquidityToken; +} + +layer signature AMMSig = { + amm : AMMInterface +} + +layer AMM : [AMMLibSig] AMMSig = { + amm = address(0x10030) <: AutomatedMarketMaker +} + +layer COMPLETE = AMM @ AMMLIB +layer L = { + o = FixedSupplyToken +} diff --git a/minic_unittests/unittests/erc20Token1.ds b/minic_unittests/unittests/erc20Token1.ds new file mode 100644 index 0000000..a7b8685 --- /dev/null +++ b/minic_unittests/unittests/erc20Token1.ds @@ -0,0 +1,404 @@ +(* TODO: add emit Events *) + +(* PROOFS: + 1. change re-defined layers + 2. balance tracks reserve + 3. increasing k -> algebra + 4. no path independence, splitting trades is strictly more expensive -> algebra, appendix D + 5. liquidity token economics is right, i.e., UNI-V2 tracks the actual share of pool --> actually, where is the trading of UNI-V2 Pair tokens? What is its value? + 6. MOD: slippage control + 7. ... cost of manipulation (hard) +*) + +object signature ERC20Interface = { + constructor : unit -> unit; + mint : address * int -> unit; + burn : address * int -> unit; + totalSupply : unit -> int; + balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object LiquidityToken : ERC20Interface { + let _totalSupply : int := 0 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + balances[msg_sender] := 100000 + + let mint (toA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply + value; + let to_bal = balances[toA] in + balances[toA] := to_bal + value + + let burn (fromA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply - value; + let from_bal = balances[fromA] in + balances[fromA] := from_bal - value + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +object signature FixedERC20Interface = { + constructor : unit -> unit; + const totalSupply : unit -> int; + const balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object FixedSupplyToken : FixedERC20Interface { + let _totalSupply : int := 100000 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + _totalSupply := 100000; + balances[msg_sender] := 100000 + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +(* this AMM supports ERC20 tokens, ETH can be incorporated as WETH for trading *) +object signature AMMInterface = { + constructor : unit -> unit; + (* swap : int * int * address -> unit; *) + + coarsed simpleSwap0 : address -> int; (* transfer token0 to amm, results number of token1 swapped *) + + (* router *) + (* swapTokensForExactTokens + swapExactTokensForTokens + addLiquidity (* why restrain the liquidity to not depreciating an asset too much? *) + removeLiquidity *) + + const getReserve0 : unit -> int; + const getReserve1 : unit -> int; + const getBalance0 : unit -> int; + + (* fund management functions *) + coarsed mint : address -> unit; (* provide liquidity to this pair *) + coarsed burn : address -> unit; (* remove liquidity from this pair *) + + coarsed skim : address -> unit; (* arbitrage to force balances to match reserves *) + coarsed sync : unit -> unit; (* force reserves to match balances *) + k : unit -> int; (* get the constant product *) + + (* oracles *) + (* we do not provide Uniswap V2 version of smoothed oracles at this moment *) + quote0 : int -> int; (* given some amount of an asset and pair reserves, returns an equivalent amount of the other asset *) + getAmountOut0 : int -> int; (* given an input amount of an asset and pair reserves, returns the maximum output amount of the other asset *) + getAmountIn0 : int -> int; (* given an output amount of an asset and pair reserves, returns a required input amount of the other asset *) +} + +(* assuming 0.3% txn fee, excluding the 0.05% additional fee in Uniswap V2 design *) +(* we do not support liquidity tokens at the moment *) +object AutomatedMarketMaker + ( liquidityToken: ERC20Interface, + erc20Token0: FixedERC20Interface, + erc20Token1: FixedERC20Interface) : AMMInterface { + + let _token0 : address := address(0xdac17f958d2ee523a2206206994597c13d831ec7) (* USDT *) + let _token1 : address := address(0x6b175474e89094c44da98b954eedeac495271d0f) (* DAI *) + let _owner : address := address(0x0) + let _reserve0 : int := 0 + let _reserve1 : int := 0 + let blockTimestampLast : int := 0 (* new oracle *) + let price0CumulativeLast : int := 0 (* new oracle *) + let price1CumulativeLast : int := 0 (* new oracle *) + let kLast : int := 0 (* reserve0 * reserve1, as of immediately after the most recent liquidity event *) + + let constructor () = + _owner := msg_sender + + let getReserve0 () = + let reserve0 = _reserve0 in + reserve0 + + let getReserve1 () = + let reserve1 = _reserve1 in + reserve1 + + let getBalance0 () = + let balance0 = erc20Token0.balanceOf(this_address) in + balance0 + + (* transfer token0 and token1 to this contract, then mint liquidity tokens *) + let mint (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0 = balance0 - reserve0 in + let amount1 = balance1 - reserve1 in + (* update reserve 0 and reserve 1 *) + let totalSupply = liquidityToken.totalSupply() in + let liquidity = if totalSupply = 0 + then + begin + liquidityToken.mint(address(0x0), 1000); (* lock the first liquidity tokens*) + amount0 * amount1 - 1000 (* in Uniswap this should be square-rooted *) + end + else + let x = amount0 * totalSupply / reserve0 in + let y = amount1 * totalSupply / reserve1 in + if x < y then x else y + in + assert (liquidity > 0); + liquidityToken.mint(toA, liquidity); + _reserve0 := balance0; + _reserve1 := balance1 + + (* need toA first transfer liquidity tokens to this contract, then operate *) + (* it does not use ERC20 style allowances, just direct transfer *) + let burn (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let liquidity = liquidityToken.balanceOf(this_address) in + let totalSupply = liquidityToken.totalSupply() in + let amount0 = liquidity * balance0 / totalSupply in + let amount1 = liquidity * balance1 / totalSupply in + assert (amount0 > 0 /\ amount1 > 0); + liquidityToken.burn(this_address, liquidity); + let success = erc20Token0.transfer(toA, amount0) in + assert (success); + let success = erc20Token1.transfer(toA, amount1) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + (* adhere to the stanford paper formalization, targeted for vanilla proof *) + let simpleSwap0 (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = balance0 - reserve0 in + let token0 = _token0 in + let token1 = _token1 in + assert (toA <> token0 /\ toA <> token1); + assert (amount0In > 0); + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let result = numerator / denominator in + let success = erc20Token1.transfer(toA, result) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1; + let resultU = result in + resultU + + (* derives from Uniswap V2 flashswap, but not flashswap since DeepSEA does not allow control-flow aggregation *) + (* first transfer tokens to this contract, then call swap to swap them *) + (* notice this does not refund, so possibly not optimal rate *) + (* let swap (amount0Out, amount1Out, toA) = + assert (amount0Out > 0 \/ amount1Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let token0 = _token0 in + let token1 = _token1 in + assert (amount0Out < reserve0 /\ amount1Out < reserve1); + assert (toA <> token0 /\ toA <> token1); + (* check the assets satisfy what is wanted *) + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + (* we can prove that balance > reserve here *) + let amount0In = helper.getAmountIn(balance0, 0, reserve0) in (* balance0 - reserve0 *) + let amount1In = helper.getAmountIn(balance1, 0, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0 - amount0Out, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1 - amount1Out, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); + + (* sending money *) + let _ = if amount0Out > 0 then + let success = erc20Token0.transfer(toA, amount0Out) in + assert (success); + () + else + () + in + let _ = if amount1Out > 0 then + let success = erc20Token1.transfer(toA, amount1Out) in + assert (success); + () + else + () + in + (* flash swap style: sent money (amount0Out, amount1Out) to msg_sender, now wait to get back *) + (* after got money back *) + (* let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = helper.getAmountIn(balance0, amount0Out, reserve0) in + let amount1In = helper.getAmountIn(balance1, amount1Out, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); *) + (* update the reserves, also update the oracle if possible *) + _reserve0 := balance0; + _reserve1 := balance1 *) + + let skim (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let skim0 = balance0 - reserve0 in + let skim1 = balance1 - reserve1 in + let success = erc20Token0.transfer(toA, skim0) in + assert (success); + let success = erc20Token1.transfer(toA, skim1) in + assert (success) + + let sync () = + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + let k () = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let resultU = reserve0 * reserve1 in + resultU + + (* Uniswap did not use Q112 library to do this computation *) + + (* given amount0 of token0, the equivalent value in token1 *) + let quote0 (amount0) = + assert (amount0 > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); (* prove that this can be dropped? *) + let resultU = amount0 * reserve1 / reserve0 in + resultU + + let getAmountOut0 (amount0In) = + assert (amount0In > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let resultU = numerator / denominator in + resultU + + let getAmountIn0 (amount0Out) = + assert (amount0Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let numerator = reserve1 * amount0Out * 1000 in + let denominator = (reserve0 - amount0Out) * 997 in + let resultU = (numerator / denominator) + 1 in + resultU +} + +layer signature AMMLibSig = { + erc20Token0 : FixedERC20Interface; + erc20Token1 : FixedERC20Interface; + liquidityToken : ERC20Interface; +} + +layer AMMLIB : [{}] AMMLibSig = { + erc20Token0 = address(0x10033) <: FixedSupplyToken; + erc20Token1 = address(0x10032) <: (clone FixedSupplyToken); (* cloned an object *) + liquidityToken = address(0x10031) <: LiquidityToken; +} + +layer signature AMMSig = { + amm : AMMInterface +} + +layer AMM : [AMMLibSig] AMMSig = { + amm = address(0x10030) <: AutomatedMarketMaker +} + +layer COMPLETE = AMM @ AMMLIB +layer L = { + o = FixedSupplyToken +} diff --git a/unittests/events.ds b/minic_unittests/unittests/events.ds similarity index 100% rename from unittests/events.ds rename to minic_unittests/unittests/events.ds diff --git a/unittests/forloop.ds b/minic_unittests/unittests/forloop.ds old mode 100644 new mode 100755 similarity index 100% rename from unittests/forloop.ds rename to minic_unittests/unittests/forloop.ds diff --git a/unittests/keccak256.ds b/minic_unittests/unittests/keccak256.ds similarity index 100% rename from unittests/keccak256.ds rename to minic_unittests/unittests/keccak256.ds diff --git a/unittests/layers1.ds b/minic_unittests/unittests/layers1.ds old mode 100644 new mode 100755 similarity index 100% rename from unittests/layers1.ds rename to minic_unittests/unittests/layers1.ds diff --git a/minic_unittests/unittests/liquidityToken.ds b/minic_unittests/unittests/liquidityToken.ds new file mode 100644 index 0000000..b25d5e2 --- /dev/null +++ b/minic_unittests/unittests/liquidityToken.ds @@ -0,0 +1,404 @@ +(* TODO: add emit Events *) + +(* PROOFS: + 1. change re-defined layers + 2. balance tracks reserve + 3. increasing k -> algebra + 4. no path independence, splitting trades is strictly more expensive -> algebra, appendix D + 5. liquidity token economics is right, i.e., UNI-V2 tracks the actual share of pool --> actually, where is the trading of UNI-V2 Pair tokens? What is its value? + 6. MOD: slippage control + 7. ... cost of manipulation (hard) +*) + +object signature ERC20Interface = { + constructor : unit -> unit; + mint : address * int -> unit; + burn : address * int -> unit; + totalSupply : unit -> int; + balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object LiquidityToken : ERC20Interface { + let _totalSupply : int := 0 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + balances[msg_sender] := 100000 + + let mint (toA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply + value; + let to_bal = balances[toA] in + balances[toA] := to_bal + value + + let burn (fromA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply - value; + let from_bal = balances[fromA] in + balances[fromA] := from_bal - value + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +object signature FixedERC20Interface = { + constructor : unit -> unit; + const totalSupply : unit -> int; + const balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object FixedSupplyToken : FixedERC20Interface { + let _totalSupply : int := 100000 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + _totalSupply := 100000; + balances[msg_sender] := 100000 + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +(* this AMM supports ERC20 tokens, ETH can be incorporated as WETH for trading *) +object signature AMMInterface = { + constructor : unit -> unit; + (* swap : int * int * address -> unit; *) + + coarsed simpleSwap0 : address -> int; (* transfer token0 to amm, results number of token1 swapped *) + + (* router *) + (* swapTokensForExactTokens + swapExactTokensForTokens + addLiquidity (* why restrain the liquidity to not depreciating an asset too much? *) + removeLiquidity *) + + const getReserve0 : unit -> int; + const getReserve1 : unit -> int; + const getBalance0 : unit -> int; + + (* fund management functions *) + coarsed mint : address -> unit; (* provide liquidity to this pair *) + coarsed burn : address -> unit; (* remove liquidity from this pair *) + + coarsed skim : address -> unit; (* arbitrage to force balances to match reserves *) + coarsed sync : unit -> unit; (* force reserves to match balances *) + k : unit -> int; (* get the constant product *) + + (* oracles *) + (* we do not provide Uniswap V2 version of smoothed oracles at this moment *) + quote0 : int -> int; (* given some amount of an asset and pair reserves, returns an equivalent amount of the other asset *) + getAmountOut0 : int -> int; (* given an input amount of an asset and pair reserves, returns the maximum output amount of the other asset *) + getAmountIn0 : int -> int; (* given an output amount of an asset and pair reserves, returns a required input amount of the other asset *) +} + +(* assuming 0.3% txn fee, excluding the 0.05% additional fee in Uniswap V2 design *) +(* we do not support liquidity tokens at the moment *) +object AutomatedMarketMaker + ( liquidityToken: ERC20Interface, + erc20Token0: FixedERC20Interface, + erc20Token1: FixedERC20Interface) : AMMInterface { + + let _token0 : address := address(0xdac17f958d2ee523a2206206994597c13d831ec7) (* USDT *) + let _token1 : address := address(0x6b175474e89094c44da98b954eedeac495271d0f) (* DAI *) + let _owner : address := address(0x0) + let _reserve0 : int := 0 + let _reserve1 : int := 0 + let blockTimestampLast : int := 0 (* new oracle *) + let price0CumulativeLast : int := 0 (* new oracle *) + let price1CumulativeLast : int := 0 (* new oracle *) + let kLast : int := 0 (* reserve0 * reserve1, as of immediately after the most recent liquidity event *) + + let constructor () = + _owner := msg_sender + + let getReserve0 () = + let reserve0 = _reserve0 in + reserve0 + + let getReserve1 () = + let reserve1 = _reserve1 in + reserve1 + + let getBalance0 () = + let balance0 = erc20Token0.balanceOf(this_address) in + balance0 + + (* transfer token0 and token1 to this contract, then mint liquidity tokens *) + let mint (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0 = balance0 - reserve0 in + let amount1 = balance1 - reserve1 in + (* update reserve 0 and reserve 1 *) + let totalSupply = liquidityToken.totalSupply() in + let liquidity = if totalSupply = 0 + then + begin + liquidityToken.mint(address(0x0), 1000); (* lock the first liquidity tokens*) + amount0 * amount1 - 1000 (* in Uniswap this should be square-rooted *) + end + else + let x = amount0 * totalSupply / reserve0 in + let y = amount1 * totalSupply / reserve1 in + if x < y then x else y + in + assert (liquidity > 0); + liquidityToken.mint(toA, liquidity); + _reserve0 := balance0; + _reserve1 := balance1 + + (* need toA first transfer liquidity tokens to this contract, then operate *) + (* it does not use ERC20 style allowances, just direct transfer *) + let burn (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let liquidity = liquidityToken.balanceOf(this_address) in + let totalSupply = liquidityToken.totalSupply() in + let amount0 = liquidity * balance0 / totalSupply in + let amount1 = liquidity * balance1 / totalSupply in + assert (amount0 > 0 /\ amount1 > 0); + liquidityToken.burn(this_address, liquidity); + let success = erc20Token0.transfer(toA, amount0) in + assert (success); + let success = erc20Token1.transfer(toA, amount1) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + (* adhere to the stanford paper formalization, targeted for vanilla proof *) + let simpleSwap0 (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = balance0 - reserve0 in + let token0 = _token0 in + let token1 = _token1 in + assert (toA <> token0 /\ toA <> token1); + assert (amount0In > 0); + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let result = numerator / denominator in + let success = erc20Token1.transfer(toA, result) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1; + let resultU = result in + resultU + + (* derives from Uniswap V2 flashswap, but not flashswap since DeepSEA does not allow control-flow aggregation *) + (* first transfer tokens to this contract, then call swap to swap them *) + (* notice this does not refund, so possibly not optimal rate *) + (* let swap (amount0Out, amount1Out, toA) = + assert (amount0Out > 0 \/ amount1Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let token0 = _token0 in + let token1 = _token1 in + assert (amount0Out < reserve0 /\ amount1Out < reserve1); + assert (toA <> token0 /\ toA <> token1); + (* check the assets satisfy what is wanted *) + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + (* we can prove that balance > reserve here *) + let amount0In = helper.getAmountIn(balance0, 0, reserve0) in (* balance0 - reserve0 *) + let amount1In = helper.getAmountIn(balance1, 0, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0 - amount0Out, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1 - amount1Out, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); + + (* sending money *) + let _ = if amount0Out > 0 then + let success = erc20Token0.transfer(toA, amount0Out) in + assert (success); + () + else + () + in + let _ = if amount1Out > 0 then + let success = erc20Token1.transfer(toA, amount1Out) in + assert (success); + () + else + () + in + (* flash swap style: sent money (amount0Out, amount1Out) to msg_sender, now wait to get back *) + (* after got money back *) + (* let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = helper.getAmountIn(balance0, amount0Out, reserve0) in + let amount1In = helper.getAmountIn(balance1, amount1Out, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); *) + (* update the reserves, also update the oracle if possible *) + _reserve0 := balance0; + _reserve1 := balance1 *) + + let skim (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let skim0 = balance0 - reserve0 in + let skim1 = balance1 - reserve1 in + let success = erc20Token0.transfer(toA, skim0) in + assert (success); + let success = erc20Token1.transfer(toA, skim1) in + assert (success) + + let sync () = + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + let k () = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let resultU = reserve0 * reserve1 in + resultU + + (* Uniswap did not use Q112 library to do this computation *) + + (* given amount0 of token0, the equivalent value in token1 *) + let quote0 (amount0) = + assert (amount0 > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); (* prove that this can be dropped? *) + let resultU = amount0 * reserve1 / reserve0 in + resultU + + let getAmountOut0 (amount0In) = + assert (amount0In > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let resultU = numerator / denominator in + resultU + + let getAmountIn0 (amount0Out) = + assert (amount0Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let numerator = reserve1 * amount0Out * 1000 in + let denominator = (reserve0 - amount0Out) * 997 in + let resultU = (numerator / denominator) + 1 in + resultU +} + +layer signature AMMLibSig = { + erc20Token0 : FixedERC20Interface; + erc20Token1 : FixedERC20Interface; + liquidityToken : ERC20Interface; +} + +layer AMMLIB : [{}] AMMLibSig = { + erc20Token0 = address(0x10033) <: FixedSupplyToken; + erc20Token1 = address(0x10032) <: (clone FixedSupplyToken); (* cloned an object *) + liquidityToken = address(0x10031) <: LiquidityToken; +} + +layer signature AMMSig = { + amm : AMMInterface +} + +layer AMM : [AMMLibSig] AMMSig = { + amm = address(0x10030) <: AutomatedMarketMaker +} + +layer COMPLETE = AMM @ AMMLIB +layer L = { + o = LiquidityToken +} diff --git a/unittests/match-int.ds b/minic_unittests/unittests/match-int.ds old mode 100644 new mode 100755 similarity index 100% rename from unittests/match-int.ds rename to minic_unittests/unittests/match-int.ds diff --git a/minic_unittests/unittests/simple_call.ds b/minic_unittests/unittests/simple_call.ds new file mode 100644 index 0000000..6b854f5 --- /dev/null +++ b/minic_unittests/unittests/simple_call.ds @@ -0,0 +1,83 @@ +object signature FixedERC20Interface = { + const getNum : unit -> int; + doubleNum : unit -> unit; + const sum : int * int -> int; +} + +object FixedSupplyToken : FixedERC20Interface { + let myNum : int := 5 + + let getNum () = + let num = myNum in + num + + let doubleNum () = + let num = myNum in + myNum := num * 2; + () + + let sum (num1, num2) = + let ret = (2 * num1) + num2 in + ret +} + +object signature AMMInterface = { + call_double : unit -> unit; + const call_get : unit -> int; + const call_sum : int * int -> int; + const stress_call : int * int -> int; +} + +object AutomatedMarketMaker (erc20Token0: FixedERC20Interface) : AMMInterface { + let call_double () = + erc20Token0.doubleNum() + + let call_get () = + let ret = erc20Token0.getNum() in + ret + + let call_sum (num1, num2) = + let ret = erc20Token0.sum(num1, num2) in + ret + + let stress_call (num1, num2) = + let temp1 = 1 in + let temp2 = 2 in + let temp3 = 3 in + let temp4 = 4 in + let temp5 = 5 in + (*let temp6 = 6 in (* breaks down *) + let temp7 = 7 in + let temp8 = 8 in + let temp9 = 9 in + let temp10 = 10 in + let temp11 = 11 in + let temp12 = 12 in + let temp13 = 13 in + let temp14 = 14 in + let temp15 = 15 in + let temp16 = 16 in + let temp16 = 16 in*) + let ret = erc20Token0.sum(num1, num2) in + ret +} + +(* CUT *) + +layer signature AMMLibSig = { + erc20Token0 : FixedERC20Interface; +} + +layer AMMLIB : [{}] AMMLibSig = { + erc20Token0 = address(0x10033) <: FixedSupplyToken; +} + +layer signature AMMSig = { + amm : AMMInterface +} + +layer AMM : [AMMLibSig] AMMSig = { + amm = address(0x10030) <: AutomatedMarketMaker +} + +layer COMPLETE = AMM @ AMMLIB diff --git a/unittests/struct.ds b/minic_unittests/unittests/struct.ds old mode 100644 new mode 100755 similarity index 100% rename from unittests/struct.ds rename to minic_unittests/unittests/struct.ds diff --git a/src/Edsger/.merlin b/src/Edsger/.merlin new file mode 100644 index 0000000..263bf33 --- /dev/null +++ b/src/Edsger/.merlin @@ -0,0 +1,111 @@ +EXCLUDE_QUERY_DIR +B /Users/zachpage/.opam/4.09.0/lib/base +B /Users/zachpage/.opam/4.09.0/lib/base/base_internalhash_types +B /Users/zachpage/.opam/4.09.0/lib/base/caml +B /Users/zachpage/.opam/4.09.0/lib/base/md5 +B /Users/zachpage/.opam/4.09.0/lib/base/shadow_stdlib +B /Users/zachpage/.opam/4.09.0/lib/base_bigstring +B /Users/zachpage/.opam/4.09.0/lib/base_quickcheck +B /Users/zachpage/.opam/4.09.0/lib/bin_prot +B /Users/zachpage/.opam/4.09.0/lib/bin_prot/shape +B /Users/zachpage/.opam/4.09.0/lib/biniou +B /Users/zachpage/.opam/4.09.0/lib/bytes +B /Users/zachpage/.opam/4.09.0/lib/core +B /Users/zachpage/.opam/4.09.0/lib/core/error_checking_mutex +B /Users/zachpage/.opam/4.09.0/lib/core_kernel +B /Users/zachpage/.opam/4.09.0/lib/core_kernel/base_for_tests +B /Users/zachpage/.opam/4.09.0/lib/core_kernel/caml_unix +B /Users/zachpage/.opam/4.09.0/lib/core_kernel/flags +B /Users/zachpage/.opam/4.09.0/lib/core_kernel/version_util +B /Users/zachpage/.opam/4.09.0/lib/cryptokit +B /Users/zachpage/.opam/4.09.0/lib/easy-format +B /Users/zachpage/.opam/4.09.0/lib/extlib +B /Users/zachpage/.opam/4.09.0/lib/fieldslib +B /Users/zachpage/.opam/4.09.0/lib/jane-street-headers +B /Users/zachpage/.opam/4.09.0/lib/ocaml +B /Users/zachpage/.opam/4.09.0/lib/ocaml/threads +B /Users/zachpage/.opam/4.09.0/lib/parsexp +B /Users/zachpage/.opam/4.09.0/lib/ppx_assert/runtime-lib +B /Users/zachpage/.opam/4.09.0/lib/ppx_bench/runtime-lib +B /Users/zachpage/.opam/4.09.0/lib/ppx_compare/runtime-lib +B /Users/zachpage/.opam/4.09.0/lib/ppx_enumerate/runtime-lib +B /Users/zachpage/.opam/4.09.0/lib/ppx_expect/collector +B /Users/zachpage/.opam/4.09.0/lib/ppx_expect/common +B /Users/zachpage/.opam/4.09.0/lib/ppx_expect/config +B /Users/zachpage/.opam/4.09.0/lib/ppx_expect/config_types +B /Users/zachpage/.opam/4.09.0/lib/ppx_hash/runtime-lib +B /Users/zachpage/.opam/4.09.0/lib/ppx_here/runtime-lib +B /Users/zachpage/.opam/4.09.0/lib/ppx_inline_test/config +B /Users/zachpage/.opam/4.09.0/lib/ppx_inline_test/runtime-lib +B /Users/zachpage/.opam/4.09.0/lib/ppx_module_timer/runtime +B /Users/zachpage/.opam/4.09.0/lib/ppx_sexp_conv/runtime-lib +B /Users/zachpage/.opam/4.09.0/lib/sexplib +B /Users/zachpage/.opam/4.09.0/lib/sexplib/unix +B /Users/zachpage/.opam/4.09.0/lib/sexplib0 +B /Users/zachpage/.opam/4.09.0/lib/spawn +B /Users/zachpage/.opam/4.09.0/lib/splittable_random +B /Users/zachpage/.opam/4.09.0/lib/stdio +B /Users/zachpage/.opam/4.09.0/lib/time_now +B /Users/zachpage/.opam/4.09.0/lib/timezone +B /Users/zachpage/.opam/4.09.0/lib/typerep +B /Users/zachpage/.opam/4.09.0/lib/variantslib +B /Users/zachpage/.opam/4.09.0/lib/yojson +B /Users/zachpage/.opam/4.09.0/lib/zarith +B ../_build/default/Edsger/.edsger.eobjs/byte +B ../_build/default/backend/extraction/.backend.objs/byte +S /Users/zachpage/.opam/4.09.0/lib/base +S /Users/zachpage/.opam/4.09.0/lib/base/base_internalhash_types +S /Users/zachpage/.opam/4.09.0/lib/base/caml +S /Users/zachpage/.opam/4.09.0/lib/base/md5 +S /Users/zachpage/.opam/4.09.0/lib/base/shadow_stdlib +S /Users/zachpage/.opam/4.09.0/lib/base_bigstring +S /Users/zachpage/.opam/4.09.0/lib/base_quickcheck +S /Users/zachpage/.opam/4.09.0/lib/bin_prot +S /Users/zachpage/.opam/4.09.0/lib/bin_prot/shape +S /Users/zachpage/.opam/4.09.0/lib/biniou +S /Users/zachpage/.opam/4.09.0/lib/bytes +S /Users/zachpage/.opam/4.09.0/lib/core +S /Users/zachpage/.opam/4.09.0/lib/core/error_checking_mutex +S /Users/zachpage/.opam/4.09.0/lib/core_kernel +S /Users/zachpage/.opam/4.09.0/lib/core_kernel/base_for_tests +S /Users/zachpage/.opam/4.09.0/lib/core_kernel/caml_unix +S /Users/zachpage/.opam/4.09.0/lib/core_kernel/flags +S /Users/zachpage/.opam/4.09.0/lib/core_kernel/version_util +S /Users/zachpage/.opam/4.09.0/lib/cryptokit +S /Users/zachpage/.opam/4.09.0/lib/easy-format +S /Users/zachpage/.opam/4.09.0/lib/extlib +S /Users/zachpage/.opam/4.09.0/lib/fieldslib +S /Users/zachpage/.opam/4.09.0/lib/jane-street-headers +S /Users/zachpage/.opam/4.09.0/lib/ocaml +S /Users/zachpage/.opam/4.09.0/lib/ocaml/threads +S /Users/zachpage/.opam/4.09.0/lib/parsexp +S /Users/zachpage/.opam/4.09.0/lib/ppx_assert/runtime-lib +S /Users/zachpage/.opam/4.09.0/lib/ppx_bench/runtime-lib +S /Users/zachpage/.opam/4.09.0/lib/ppx_compare/runtime-lib +S /Users/zachpage/.opam/4.09.0/lib/ppx_enumerate/runtime-lib +S /Users/zachpage/.opam/4.09.0/lib/ppx_expect/collector +S /Users/zachpage/.opam/4.09.0/lib/ppx_expect/common +S /Users/zachpage/.opam/4.09.0/lib/ppx_expect/config +S /Users/zachpage/.opam/4.09.0/lib/ppx_expect/config_types +S /Users/zachpage/.opam/4.09.0/lib/ppx_hash/runtime-lib +S /Users/zachpage/.opam/4.09.0/lib/ppx_here/runtime-lib +S /Users/zachpage/.opam/4.09.0/lib/ppx_inline_test/config +S /Users/zachpage/.opam/4.09.0/lib/ppx_inline_test/runtime-lib +S /Users/zachpage/.opam/4.09.0/lib/ppx_module_timer/runtime +S /Users/zachpage/.opam/4.09.0/lib/ppx_sexp_conv/runtime-lib +S /Users/zachpage/.opam/4.09.0/lib/sexplib +S /Users/zachpage/.opam/4.09.0/lib/sexplib/unix +S /Users/zachpage/.opam/4.09.0/lib/sexplib0 +S /Users/zachpage/.opam/4.09.0/lib/spawn +S /Users/zachpage/.opam/4.09.0/lib/splittable_random +S /Users/zachpage/.opam/4.09.0/lib/stdio +S /Users/zachpage/.opam/4.09.0/lib/time_now +S /Users/zachpage/.opam/4.09.0/lib/timezone +S /Users/zachpage/.opam/4.09.0/lib/typerep +S /Users/zachpage/.opam/4.09.0/lib/variantslib +S /Users/zachpage/.opam/4.09.0/lib/yojson +S /Users/zachpage/.opam/4.09.0/lib/zarith +S . +S ../backend/extraction +FLG -pp '/Users/zachpage/.opam/4.09.0/bin/cppo -n -V OCAML:4.09.0' +FLG -w -40 diff --git a/src/Edsger/ast.ml b/src/Edsger/ast.ml index c92d23b..3c0a9cc 100755 --- a/src/Edsger/ast.ml +++ b/src/Edsger/ast.ml @@ -211,7 +211,8 @@ and a_command_desc = | ACyield of a_rexpr | AClet of tmp_id_t * ident * a_command * a_command | ACsequence of a_command * a_command - | ACcall of ident * ident * a_rexpr list + (* ACcall: layer signature name * method name * args * value* gas*) + | ACcall of ident * ident * a_rexpr list * a_rexpr option * a_rexpr option | ACcond of a_rexpr * a_command * a_command | ACfor of tmp_id_t * ident * a_rexpr * tmp_id_t * a_rexpr * a_command (* FOR (*NUM*) i = 0 TO (*HOLDER*) n DO s.run i *) @@ -595,7 +596,7 @@ let rec string_of_a_command c = match c.aCmdDesc with "\nin " ^ string_of_a_command c2 | ACsequence (c1, c2) -> string_of_a_command c1 ^ ";\n" ^ string_of_a_command c2 - | ACcall (s, f, es) -> "call " ^ s ^ "." ^ f ^ " (" ^ + | ACcall (s, f, es, _, _) -> "call " ^ s ^ "." ^ f ^ " (" ^ String.concat ", " (List.map string_of_a_rexpr es) ^ ")" | ACcond (e, c1, c2) -> "if " ^ string_of_a_rexpr e ^ "\n then " ^ string_of_a_command c1 ^ diff --git a/src/Edsger/ast.mli b/src/Edsger/ast.mli index 20f30e9..aef8893 100755 --- a/src/Edsger/ast.mli +++ b/src/Edsger/ast.mli @@ -191,7 +191,7 @@ and a_command_desc = | ACyield of a_rexpr | AClet of tmp_id_t * ident * a_command * a_command | ACsequence of a_command * a_command - | ACcall of ident * ident * a_rexpr list + | ACcall of ident * ident * a_rexpr list * a_rexpr option * a_rexpr option | ACcond of a_rexpr * a_command * a_command | ACfor of tmp_id_t * ident * a_rexpr * tmp_id_t * a_rexpr * a_command (* FOR (*NUM*) i = 0 TO (*HOLDER*) n DO s.run i *) diff --git a/src/Edsger/config.h b/src/Edsger/config.h index 4a6c98a..eba4114 100644 --- a/src/Edsger/config.h +++ b/src/Edsger/config.h @@ -1 +1,2 @@ -(* #define ANT *) +#define REDACTED +#define ANT diff --git a/src/Edsger/coqgen.ml b/src/Edsger/coqgen.ml index af17066..95c0dfb 100755 --- a/src/Edsger/coqgen.ml +++ b/src/Edsger/coqgen.ml @@ -619,7 +619,7 @@ let output_ft_cond out ind f t = output_char out ')' | ATmapping _ -> coqgen_warning "Don't know how to output_ft_cond of a mapping"; - output_string out "(* TODO: implement this in output_ft_cond. *)" + output_string out "True (* TODO: implement this in output_ft_cond. *)" | ATlist _ -> coqgen_fatal_error __LOC__ "output_ft_cond" ("External type " ^ string_of_a_type false t ^ " marked as provable") @@ -732,7 +732,7 @@ let output_valid_ft_cond out ind f t = output_char out ')' | ATmapping _ -> coqgen_warning "Don't know how to output_valid_ft_cond of a mapping"; - output_string out "(* TODO: implement this in output_ft_cond. *)" + output_string out "True (* TODO: implement this in output_ft_cond. *)" | ATlist t' -> coqgen_fatal_error __LOC__ "output_valid_ft_cond" ("List type " ^ string_of_a_type false t ^ " marked as provable") @@ -1070,7 +1070,6 @@ let output_record_hyper_field_impl out struct_type_pair_ident s f ft output_string out ("Global Instance " ^ struct_type_pair_ident ^ "_" ^ f ^ "_field_impl : HyperFieldImpl " ^ struct_type_pair_ident ^ " " ^ ft.aTypePairIdent ^ " " ^ struct_field_name_to_ident s f ^ " := { - Hfield_offset := " ^ string_of_int offset ^ "; Hfield_get s := "); output_get out "s"; (* fst s *) output_string out ";\n Hfield_set v s := "; @@ -1091,7 +1090,6 @@ let output_twobranch_hyper_field_impl out t s cs f ft cf output_string out ("Global Instance " ^ t.aTypePairIdent ^ "_" ^ f ^ "_field_impl : HyperFieldImpl " ^ t.aTypePairIdent ^ " " ^ ft.aTypePairIdent ^ " " ^ struct_field_name_to_ident cs cf ^ " := { - Hfield_offset := " ^ string_of_int offset ^ "; Hfield_get s := match s with \n"); (* TODO: generate fresh names instead of just "s" and "v". *) output_string out (" | " ^ valid_branch.aTypeConstrName ^ " "); @@ -1128,7 +1126,7 @@ let output_record_hyper_field out struct_t s f t = "_" ^ f ^ "_field : HyperField " ^ struct_t.aTypePairIdent ^ " " ^ t.aTypePairIdent ^ " " ^ struct_field_name_to_ident s f ^ ". -Proof. solve_record_type_hyper_field. Qed.\n") +Proof. Admitted. \n") let output_twobranch_hyper_field out struct_t s cs f cf t = if struct_t.aTypeProvable then @@ -1504,7 +1502,7 @@ let rec output_rexpr out ind e = match e.aRexprDesc with ("Internal error, encountered unknown builtin \""^other^"\".") let rec output_lexpr out obj ind e = match e.aLexprDesc with - | AEglob i -> output_string out ("(LCvar " ^ obj ^ "_" ^ i ^ "_global)") + | AEglob i -> output_string out ("(LCvar " ^ obj ^ "_" ^ i ^ "_var)") | AEfield (e', f) -> output_string out ("(LCfield " ^ e.aLexprType.aTypePairIdent ^ " " ^ begin match e'.aLexprType.aTypeDesc with @@ -1517,7 +1515,11 @@ let rec output_lexpr out obj ind e = match e.aLexprDesc with output_lexpr out obj (" " ^ ind) e'; output_char out ')' | AEindex (e', idx) -> - output_string out ("(LCindex " ^ e.aLexprType.aTypePairIdent ^ "\n" ^ ind); + (* print_endline (string_of_a_type true e'.aLexprType); *) + let lc_type = match e'.aLexprType.aTypeDesc with + | ATmapping _ -> "LChash" + | _ -> "LCindex" in + output_string out ("(" ^ lc_type ^ " " ^ e.aLexprType.aTypePairIdent ^ "\n" ^ ind); output_lexpr out obj (" " ^ ind) e'; output_string out ("\n" ^ ind); output_rexpr out (" " ^ ind) idx; @@ -1913,7 +1915,7 @@ let output_command env out base_layer obj method_full_name = output_string out ("\n" ^ ind); output (" " ^ ind) pure (2::path) c2; output_char out ')' - | ACcall (s, f, es) -> + | ACcall (s, f, es, _, _) -> let o = (try (List.assoc s base_layer.aLayerAllObjects) with Not_found -> raise (PrimitiveNotFound s); Typecheck.dummy_object s) in (match o.aObjectAddress with @@ -2499,7 +2501,7 @@ Section OBJECT_" ^ i ^ "_DEFINITION. output_string out (" Definition " ^ variable_coq_name ^ " := {| ltype_tp_marker := " ^ f.aObjectFieldType.aTypePairIdent ^ "; - ltype_ident := Values.Iident " ^ ( + ltype_ident := HighValues.Field HighValues.Global " ^ ( if f.aObjectFieldIsLogical then string_of_ident 99 else @@ -2586,7 +2588,7 @@ Section OBJECT_" ^ i ^ "_DEFINITION. PRIMpure := " ^ string_of_bool is_pure ^ "; PRIMargt_marker := " ^ method_full_name ^ ".(FC_params); PRIMret_marker := " ^ method_full_name ^ ".(FC_returns); - (* PRIMcond := " ^ method_full_name ^ "_spec_cond; *) + PRIMcond := fun _ _ _ => True; (* PRIMsem := " ^ method_full_name ^ "_spec_hlist; *) PRIMsem_opt := " ^ method_full_name ^ "_spec_hlist_opt |}. @@ -3455,6 +3457,7 @@ let gen_layer env i l = "Require Import DeepSpec.lib.Monad.OptionMonad.\n" ^ "Require Import DeepSpec.lib.Monad.MonadZero.\n" ^ "Require Import DeepSpec.core.SynthesisStmt.\n" ^ + "Require Import DeepSpec.core.SynthesisFunc.\n" ^ "Require Import backend.MachineModel.\n" ^ "Existing Instance MonadState_DS.\n" ^ "Existing Instance MonadZero_DS.\n" ^ @@ -3465,9 +3468,9 @@ let gen_layer env i l = "") file_class_Layers in let _ = output_string out "\n -(*Context {memModelOps : MemoryModelOps mem}.*) +Context {memModelOps : MemoryModelOps mem}. Instance GlobalLayerSpec : LayerSpecClass := { - (*memModelOps := memModelOps;*) + memModelOps := memModelOps; GetHighData := global_abstract_data_type }. \n" in @@ -4252,16 +4255,17 @@ let do_gen_layer_code env i l = m.aMethodType.aMethodKind <> MKconstghost ) o.aObjectMethods then begin let proofs, prf_tbl = -#ifdef REDACTED - open_out "/dev/null", Hashtbl.create 100 in -#else new_incremental_file env ("Obj" ^ o.aObjectName ^ "CodeProofs") - ("Require Import liblayers.compcertx.MakeProgram.\n" ^ + ( +#ifndef REDACTED + "Require Import liblayers.compcertx.MakeProgram.\n" ^ "Require Import liblayers.compcertx.MemWithData.\n\n" ^ +#else "Require Import " ^ env.project_name ^ ".Layer" ^ l.aLayerName ^ ".\n") (fun out -> output_string out ( "Existing Instance GlobalLayerSpec. -Existing Instances " ^ l.aLayerName ^ "_overlay_spec.\n")) +Existing Instances " ^ l.aLayerName ^ "_overlay_spec.\n\n" ^ + "Context {memModelOps : MemoryModelOps mem}.\n")) (Str.regexp ("\\bLemma +" ^ o.aObjectName ^ "_\\(.+\\)_spec_requires_kernel_mode\\b")) (* TODO: fix this. *) file_class_ObjAux in #endif @@ -4451,8 +4455,7 @@ Lemma " ^ method_full_name ^ "_vc" ^ args ^ " me d :"); synth_func_cond " ^ method_full_name ^ " " ^ method_full_name ^ "_wf " ^ args ^ " me d. Proof. - admit. -Qed. +Admitted. Lemma " ^ method_full_name ^ "_oblg" ^ args ^ " me d :"); for i = 0 to nargs - 1 do @@ -4465,8 +4468,7 @@ Lemma " ^ method_full_name ^ "_oblg" ^ args ^ " me d :"); synth_func_obligation " ^ method_full_name ^ " " ^ method_full_name ^ "_wf " ^ args ^ " me d. Proof. - admit. -Qed.\n") +Admitted.\n") end; (* [if] requires kernel mode proof not exists *) output_string code (" @@ -4561,6 +4563,8 @@ let gen_global_abstract_data_type env final_layer fileDeclarations = function me_coinbase := me_coinbase me; me_timestamp := me_timestamp me; me_number := me_number me; + me_chainid := me_chainid me; + me_selfbalance := me_selfbalance me; me_balance := me_balance me; me_blockhash := me_blockhash me; me_transfer := me_transfer me; @@ -4684,7 +4688,7 @@ let gen_global_abstract_data_type env final_layer fileDeclarations = function let gen_coqProj env fileDeclarations = let stream = open_out (env.project_name ^ "/_CoqProject") in output_string stream ( - "-R ../../.. DeepSpec\n" ^ + "-R ../../../src DeepSpec\n" ^ "-R . " ^ env.project_name ^ "\n" ^ "./EdsgerIdents.v\n" ^ "./DataTypes.v\n" ^ @@ -4695,10 +4699,18 @@ let gen_coqProj env fileDeclarations = | i, ADlayer l -> output_string stream ("./Layer" ^ i ^ ".v\n") | _, _ -> () ) fileDeclarations; + List.iter (function + | i, ADlayer l -> List.iter (fun (_, o) -> + output_string stream ("./Obj" ^ o.aObjectName ^ "CodeProofs.v\n")) + l.aLayerFreshObjects + | _, _ -> () + ) fileDeclarations; +#ifndef REDACTED List.iter (function | i, ADlayer l -> output_string stream ("./LSrc" ^ i ^ ".v\n") | _, _ -> () ) fileDeclarations; +#endif close_out stream let gen_extract_make env fileDeclarations = @@ -5105,9 +5117,8 @@ let coqgen filename ast = | Some l -> gen_global_abstract_data_type env l ast.aFileDeclarations ast.aFileGlobalAbstractDataType end; gen_linksource env ast.aFileDeclarations; - gen_prf env ast.aFileDeclarations; -#ifndef REDACTED gen_coqProj env ast.aFileDeclarations; +#ifndef REDACTED gen_extract_make env ast.aFileDeclarations; #endif delete_coqgen_env env diff --git a/src/Edsger/dune b/src/Edsger/dune index 48cb84d..6f7797f 100644 --- a/src/Edsger/dune +++ b/src/Edsger/dune @@ -2,7 +2,7 @@ (name edsger) (modes byte exe) ; The line below is a hack to build a MacOS binary without libgmp.so. - (link_flags (-cclib -Wl,/usr/local/lib/libgmp.a)) + ; (link_flags (-cclib -Wl,/usr/local/lib/libgmp.a)) (libraries backend cryptokit yojson extlib str) (preprocessor_deps config.h) (preprocess (action (run %{bin:cppo} -n -V OCAML:%{ocaml_version} %{input-file}))) diff --git a/src/Edsger/edsger.ml b/src/Edsger/edsger.ml index e1cad33..494a8c5 100755 --- a/src/Edsger/edsger.ml +++ b/src/Edsger/edsger.ml @@ -13,11 +13,11 @@ let deepsea_basename = "DeepSEA/AntChain-EVM" #else let deepsea_basename = "DeepSEA/EVM" #endif -let deepsea_version_num = "1.2.0" +let deepsea_version_num = "1.3.0" let deepsea_version = deepsea_basename ^ " " ^ deepsea_version_num -type mode = ABI | BYTECODE | BYTECODE_RUNTIME | COMBINED_JSON | ASSEMBLY | MINIC | MINIC_VERBOSE | COQ | EWASM | EWASM_RUNTIME | METADATA | UCLID | CGRAPH | CLASH +type mode = ABI | BYTECODE | BYTECODE_RUNTIME | COMBINED_JSON | ASSEMBLY | ASSEMBLY_RUNTIME | MINIC | MINIC_VERBOSE | COQ | EWASM | EWASM_RUNTIME | METADATA | UCLID | CGRAPH | CLASH let string_of_token = function | ARRAY -> "ARRAY" @@ -195,7 +195,9 @@ let main argv = | "abi" -> ABI | "combined-json" -> COMBINED_JSON | "assembly" -> ASSEMBLY + | "assembly-runtime" -> ASSEMBLY_RUNTIME | "minic" -> MINIC + | "minic-verbose" -> MINIC_VERBOSE | "coq" -> COQ | "ewasm" -> EWASM | "uclid" -> UCLID @@ -210,8 +212,8 @@ let main argv = let parse_structure = try Parser.file Lexer.token - (*fun buf -> let t = Lexer.token buf - in print_endline ("TOK: " ^ string_of_token t); t*) + (* (fun buf -> let t = Lexer.token buf *) + (* in print_endline ("TOK: " ^ string_of_token t); t) *) buf with Failure _ | Parser.Error -> @@ -232,8 +234,7 @@ let main argv = else let abi = abigen ast_structure in match mode_flag with - (* COQ -> Coqgen.coqgen filename ast_structure *) - | COQ -> print_endline "Coq output is not supported in this preview release."; exit 1 + | COQ -> Coqgen.coqgen filename ast_structure | ABI -> print_endline abi | METADATA -> print_endline (metadata deepsea_version_num argv.(0) filename abi) | _ -> @@ -277,7 +278,8 @@ let main argv = print_endline (Backend.ASM.mnemonics_clash cg')) cgd in ()) | BYTECODE -> print_endline (bytecode false ge) | BYTECODE_RUNTIME -> print_endline (bytecode true ge) - | ASSEMBLY -> print_endline (assembly ge) + | ASSEMBLY -> print_endline (assembly false ge) + | ASSEMBLY_RUNTIME -> print_endline (assembly true ge) | COMBINED_JSON -> let asm, _ , programsize = get_bytecode_params ge in print_endline (combined_json filename abi (Backend.ASM.assemble asm)) diff --git a/src/Edsger/make_parser.sh b/src/Edsger/make_parser.sh index bf0c592..7ac400a 100755 --- a/src/Edsger/make_parser.sh +++ b/src/Edsger/make_parser.sh @@ -1,6 +1,7 @@ #!/bin/bash -Edsger=`dirname $0` +root=$(git rev-parse --show-toplevel) +Edsger=$root/Edsger source_file=$(mktemp) processed=$(mktemp) diff --git a/src/Edsger/minicgen.ml b/src/Edsger/minicgen.ml index d306a2a..8625596 100755 --- a/src/Edsger/minicgen.ml +++ b/src/Edsger/minicgen.ml @@ -3,6 +3,7 @@ open Astcommon open Abi open Backend + open BinNumsExt open StmtMiniC open ExpMiniC module D = Datatypes @@ -608,18 +609,53 @@ and gen_cmd underlay obj pure c dest = | ACsequence (c1, c2) -> Ssequence (gen_cmd underlay obj pure c1 dest, gen_cmd underlay obj pure c2 dest) - | ACcall (s, f, es) -> - let retval_dest = begin match c.aCmdType.aTypeCtype with - | ACtvoid -> Backend.Datatypes.None - | _ -> Backend.Datatypes.Some (positive_of_int dest) - end in - let args = begin match es with - | [e] when e.aRexprType.aTypeDesc = ATbuiltin Tunit -> [] - | _ -> List.map gen_rexpr es - end in - Scall (retval_dest, - backend_ident_of_primitive underlay s f, - coqlist_of_list args) + | ACcall (s, f, es, v, g) -> + let o = (match underlay with + | ALontop l -> (try Some (List.assoc s l.aLayerAllObjects) + with Not_found -> None) + | _ -> None) + in + let retval_dest = begin match c.aCmdType.aTypeCtype with + | ACtvoid -> Backend.Datatypes.None + | _ -> Backend.Datatypes.Some (positive_of_int dest) + end in + let args = begin match es with + | [e] when e.aRexprType.aTypeDesc = ATbuiltin Tunit -> [] + | _ -> List.map gen_rexpr es + end in + let address_signature = (match o with + | None -> None + | Some o' -> + let methd = List.find (fun i -> i.aMethodName = f) o'.aObjectMethods in + let signature = coq_Z_of_int (function_selector_intval_of_method methd) in + match o'.aObjectAddress with + | None -> None + | Some address -> Some (address, signature) + ) in + (match address_signature with + | None -> if Option.is_some v || Option.is_some g then + raise (Failure "Value/gas specified for interal method call") else + Scall (retval_dest, + backend_ident_of_primitive underlay s f, + coqlist_of_list args) + | Some (address, signature) -> + let v' = (match v with + | None -> Econst_int256 (coq_Z_of_int 0, Tint (I256, Unsigned)) + | Some v'' -> gen_rexpr v'' + ) in + let g' = (match g with + | None -> D.None + | Some g'' -> D.Some (gen_rexpr g'') + ) in + Scallmethod ( + Econst_int256 (coq_Z_of_Z (Z.of_string address), Tint (I256, Unsigned)), + coqlist_of_list (match retval_dest with None -> [] | Some r -> [r]), (* only one return value allowed for now *) + signature, (* method siguature *) + v', + g', + coqlist_of_list args + ) + ) | ACcond (e, c_then, c_else) -> Sifthenelse (gen_rexpr e, gen_cmd underlay obj pure c_then dest, @@ -857,7 +893,7 @@ let rec string_of_statement = function | Sreturn None -> "Sreturn None" | Sreturn (Some id) -> ("Sreturn Some("^string_of_int (int_of_positive id)^")") | Stransfer (e1,e2) -> "Stransfer ("^string_of_expr e1 ^","^ string_of_expr e2 ^")" - | Scallmethod (e1,ids,z,e,es) -> "Scallmethod TODO" + | Scallmethod (e1,ids,z,e,eo,es) -> "Scallmethod TODO" | Slog _ -> "Slog TODO" | Srevert -> "Srevert" @@ -920,7 +956,7 @@ let rec gen_cmd_locals c dest = (n, c1.aCmdType.aTypeCtype) :: gen_cmd_locals c1 n @ gen_cmd_locals c2 dest | ACsequence (c1, c2) -> gen_cmd_locals c1 dest @ gen_cmd_locals c2 dest - | ACcall (s, f, es) -> [] + | ACcall (s, f, es, _, _) -> [] | ACcond (e, c1, c2) -> gen_cmd_locals c1 dest @ gen_cmd_locals c2 dest | ACfor (n_iter, i, e1, n_end, e2, c, None) -> diff --git a/src/Edsger/parser.ml b/src/Edsger/parser.ml new file mode 100644 index 0000000..7bbe663 --- /dev/null +++ b/src/Edsger/parser.ml @@ -0,0 +1,15392 @@ + +module MenhirBasics = struct + + exception Error + + type token = + | XOR + | WITH + | UNEQUAL + | UINT of ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) + | TYPE + | TRUSTED + | TO + | THEN + | STRING of ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 23 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) + | STAR + | SLASH + | SIGNATURE + | SHR + | SHL + | SEMICOLON + | RPAREN + | REFINED + | RBRACKET + | RBRACE + | PLUS + | OBJECT + | MOD + | MINUS + | MATCH + | MAPPING + | LPAREN + | LOGICAL + | LIST + | LET + | LESSEQ + | LESS + | LBRACKET + | LBRACE + | LAYER + | INT of ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 53 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) + | INDEXED + | IN + | IF + | IDENTITY + | IDENT of ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 62 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) + | GREATEREQ + | GREATER + | GHOST + | FOR + | FOLD + | FIRST + | FAIL + | EXTERNAL + | EVENT + | EQUAL + | EOF + | END + | EMIT + | ELSE + | DOUBLEARROW + | DOT + | DO + | DISJUNCTION + | DENY + | CONSTRUCTOR + | CONST + | CONJUNCTION + | COMMA + | COLONLESS + | COLONGREATER + | COLONCOLON + | COLON + | CLONE + | BITNOT + | BITAND + | BEGIN + | BARBAR + | BAR + | BANG + | AT + | ASSIGN + | ASSERT + | ARROW + | ARRAY + +end + +include MenhirBasics + +let _eRR = + MenhirBasics.Error + +type _menhir_env = { + _menhir_lexer: Lexing.lexbuf -> token; + _menhir_lexbuf: Lexing.lexbuf; + _menhir_token: token; + mutable _menhir_error: bool +} + +and _menhir_state = + | MenhirState507 + | MenhirState503 + | MenhirState501 + | MenhirState499 + | MenhirState497 + | MenhirState496 + | MenhirState495 + | MenhirState494 + | MenhirState491 + | MenhirState489 + | MenhirState485 + | MenhirState481 + | MenhirState480 + | MenhirState479 + | MenhirState477 + | MenhirState476 + | MenhirState475 + | MenhirState474 + | MenhirState473 + | MenhirState471 + | MenhirState468 + | MenhirState464 + | MenhirState458 + | MenhirState455 + | MenhirState452 + | MenhirState445 + | MenhirState439 + | MenhirState433 + | MenhirState432 + | MenhirState427 + | MenhirState426 + | MenhirState420 + | MenhirState419 + | MenhirState412 + | MenhirState411 + | MenhirState410 + | MenhirState409 + | MenhirState407 + | MenhirState402 + | MenhirState401 + | MenhirState399 + | MenhirState396 + | MenhirState393 + | MenhirState391 + | MenhirState387 + | MenhirState385 + | MenhirState376 + | MenhirState375 + | MenhirState371 + | MenhirState367 + | MenhirState361 + | MenhirState355 + | MenhirState354 + | MenhirState352 + | MenhirState351 + | MenhirState350 + | MenhirState346 + | MenhirState343 + | MenhirState342 + | MenhirState341 + | MenhirState339 + | MenhirState332 + | MenhirState331 + | MenhirState329 + | MenhirState326 + | MenhirState325 + | MenhirState317 + | MenhirState312 + | MenhirState304 + | MenhirState302 + | MenhirState294 + | MenhirState293 + | MenhirState292 + | MenhirState291 + | MenhirState287 + | MenhirState286 + | MenhirState283 + | MenhirState282 + | MenhirState281 + | MenhirState280 + | MenhirState273 + | MenhirState272 + | MenhirState266 + | MenhirState262 + | MenhirState261 + | MenhirState259 + | MenhirState256 + | MenhirState254 + | MenhirState248 + | MenhirState247 + | MenhirState246 + | MenhirState245 + | MenhirState243 + | MenhirState242 + | MenhirState240 + | MenhirState239 + | MenhirState238 + | MenhirState237 + | MenhirState235 + | MenhirState233 + | MenhirState232 + | MenhirState230 + | MenhirState228 + | MenhirState224 + | MenhirState223 + | MenhirState222 + | MenhirState220 + | MenhirState211 + | MenhirState205 + | MenhirState201 + | MenhirState193 + | MenhirState190 + | MenhirState187 + | MenhirState183 + | MenhirState178 + | MenhirState176 + | MenhirState174 + | MenhirState170 + | MenhirState162 + | MenhirState160 + | MenhirState159 + | MenhirState158 + | MenhirState156 + | MenhirState154 + | MenhirState152 + | MenhirState150 + | MenhirState149 + | MenhirState147 + | MenhirState146 + | MenhirState144 + | MenhirState143 + | MenhirState140 + | MenhirState138 + | MenhirState137 + | MenhirState135 + | MenhirState134 + | MenhirState132 + | MenhirState130 + | MenhirState127 + | MenhirState125 + | MenhirState124 + | MenhirState121 + | MenhirState120 + | MenhirState118 + | MenhirState117 + | MenhirState114 + | MenhirState111 + | MenhirState109 + | MenhirState107 + | MenhirState102 + | MenhirState96 + | MenhirState95 + | MenhirState89 + | MenhirState87 + | MenhirState85 + | MenhirState83 + | MenhirState81 + | MenhirState79 + | MenhirState77 + | MenhirState74 + | MenhirState72 + | MenhirState70 + | MenhirState68 + | MenhirState66 + | MenhirState64 + | MenhirState62 + | MenhirState60 + | MenhirState58 + | MenhirState56 + | MenhirState54 + | MenhirState52 + | MenhirState51 + | MenhirState49 + | MenhirState47 + | MenhirState46 + | MenhirState37 + | MenhirState35 + | MenhirState33 + | MenhirState32 + | MenhirState30 + | MenhirState20 + | MenhirState11 + | MenhirState8 + | MenhirState6 + | MenhirState4 + +# 2 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + +open Astcommon +open Parsetree + +let make_loc (startpos, endpos) = { + loc_start = startpos; + loc_end = endpos; +} + +let mkexp_ ~loc d = + { p_expression_desc = d; p_expression_loc = make_loc loc; } +let mkcmd ~loc d = + { p_command_desc = d; p_command_loc = make_loc loc; } +let mkfotyp ~loc d = + { p_type_FO_desc = d; p_type_FO_loc = make_loc loc; } +let mksig ~loc d = + { p_signature_desc = d; p_signature_loc = make_loc loc; } +let mklayersign ~loc d = + { p_layer_signature_desc = d; p_layer_signature_loc = make_loc loc; } +let mkobjtyp ~loc objb objs = + { pObjectBase = objb; pObjectSignature = objs; pObjectTypLoc = make_loc loc; } +let mklayertyp ~loc plbase plsig = + { pLayerBase = plbase; pLayerSignature = plsig; pLayerLoc = make_loc loc } +let mkmethdef ~loc args rtyp kind body annos = + { pMethodArguments = args; pMethodReturnType = rtyp; pMethodKind = kind; pMethodBody = body; pMethodAnnotations = annos; pMethodLoc = make_loc loc; } +let mkobjconstr ~loc objtyp objk objfs objms = + { pObjType = objtyp; pObjKind = objk; pObjFields = objfs; pObjMethods = objms; pObjLoc = make_loc loc; } +let mkobj ~loc d = + { p_object_desc = d; p_object_loc = make_loc loc; } +let mkobjdef ~loc objtyp obj = + { pObjectType = objtyp; pObjectDesc = obj; pObjectLoc = make_loc loc; } +let mkobjinst ~loc objinstdesc = + { p_object_inst_desc = objinstdesc; p_object_inst_loc = make_loc loc; } +let mkprop ~loc d = + { p_proposition_desc = d; p_proposition_loc = make_loc loc;} +let mklayer ~loc d = + { p_layer_desc = d; p_layer_loc = make_loc loc; } +let mklayerdef ~loc layer invar annos = + { pLayerDesc = layer; pLayerInvariant = invar; pLayerAnnotations = annos; pLayerLoc = make_loc loc} +let mkdecl ~loc d = + { p_declaration_desc = d; p_declaration_loc = make_loc loc; } + +(* let mkanno ~loc d = + { p_annotation_desc = d; p_annotation_loc = make_loc loc; } *) + +let builtin_type_table = + let tbl = Hashtbl.create 20 in + List.iter (fun (key, data) -> Hashtbl.add tbl key data) [ + "unit", (PTbuiltin Tunit); + "bool", (PTbuiltin Tbool); + "int", (PTbuiltin Tint); + "uint", (PTbuiltin Tuint); + "hashvalue", (PTbuiltin Thashvalue); + "globalpointer", (PTbuiltin Tglobalpointer) (*; + "val", PTbuiltin Tval; + "flatmem", PTbuiltin Tflatmem*) + ]; + tbl + +let constant_table = + let tbl = Hashtbl.create 20 in + List.iter (fun (key, data) -> Hashtbl.add tbl key data) [ + "true", PEconstant (CONbool true); + "false", PEconstant (CONbool false); + "null_hash", PEconstant CONhashvalue; + (*"array_init", PEconstant CONarray_init;*) + "GLOBUndef", PEconstant CONglobalpointer_undef (*; + "Vundef", PEconstant CONval_undef; + "flatmem_empty", PEconstant CONflatmem_empty*) + ]; + tbl + +let either_to_double_list = function + | `Left a -> [a], [] + | `Right b -> [], [b] +let cons_either_double_list e (l1, l2) = match e with + | `Left a -> a :: l1, l2 + | `Right b -> l1, b :: l2 + +# 385 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + +let rec _menhir_goto_match_clauses : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_match_clauses -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv2011 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * _menhir_state * 'tv_opt_bar) * _menhir_state * 'tv_match_clauses) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BAR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv2003 * _menhir_state * 'tv_match_clauses) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run193 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState205 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACKET -> + _menhir_run112 _menhir_env (Obj.magic _menhir_stack) MenhirState205 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState205) : 'freshtv2004) + | END -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv2007 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * _menhir_state * 'tv_opt_bar) * _menhir_state * 'tv_match_clauses) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv2005 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * _menhir_state * 'tv_opt_bar) * _menhir_state * 'tv_match_clauses) = Obj.magic _menhir_stack in + let (_endpos__6_ : Lexing.position) = _endpos in + ((let ((((_menhir_stack, _menhir_s, _startpos__1_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_), _, (_4 : 'tv_opt_bar)), _, (_5 : 'tv_match_clauses)) = _menhir_stack in + let _6 = () in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__6_ in + let _v : 'tv_command_core = let _endpos = _endpos__6_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 552 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( (PCmatch (mkexp_ ~loc:_sloc e.p_expression_desc, List.rev _5)) ) +# 429 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv2006)) : 'freshtv2008) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv2009 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * _menhir_state * 'tv_opt_bar) * _menhir_state * 'tv_match_clauses) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv2010)) : 'freshtv2012) + +and _menhir_goto_match_clause : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_match_clause -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + match _menhir_s with + | MenhirState205 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1997 * _menhir_state * 'tv_match_clauses)) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_match_clause) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1995 * _menhir_state * 'tv_match_clauses)) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_match_clause) : 'tv_match_clause) = _v in + ((let (_menhir_stack, _menhir_s, (_1 : 'tv_match_clauses)) = _menhir_stack in + let _2 = () in + let _v : 'tv_match_clauses = +# 608 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 457 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_match_clauses _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1996)) : 'freshtv1998) + | MenhirState111 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv2001) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_match_clause) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1999) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_match_clause) : 'tv_match_clause) = _v in + ((let _v : 'tv_match_clauses = +# 607 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 472 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_match_clauses _menhir_env _menhir_stack _menhir_s _v) : 'freshtv2000)) : 'freshtv2002) + | _ -> + _menhir_fail () + +and _menhir_goto_proposition : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_proposition -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v -> + match _menhir_s with + | MenhirState387 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1989 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_proposition) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1987 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) = Obj.magic _menhir_stack in + let (_endpos__5_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_5 : 'tv_proposition) : 'tv_proposition) = _v in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_layer_expression), _startpos__1_), _endpos__3_, _, (_3 : 'tv_layer_expression), _startpos__3_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__5_ in + let _v : 'tv_layer_expression = let _endpos = _endpos__5_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 768 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayer ~loc:_sloc (PLrefine (_1, _3, _5)) ) +# 503 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1988)) : 'freshtv1990) + | MenhirState396 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1993 * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_proposition) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1991 * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_2 : 'tv_proposition) : 'tv_proposition) = _v in + ((let (_menhir_stack, _startpos__1_) = _menhir_stack in + let _1 = () in + let _endpos = _endpos__2_ in + let _v : 'tv_layer_invariant_annotation = +# 757 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( Some _2 ) +# 523 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_invariant_annotation _menhir_env _menhir_stack _endpos _v) : 'freshtv1992)) : 'freshtv1994) + | _ -> + _menhir_fail () + +and _menhir_goto_commands : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_commands -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v) in + match _menhir_s with + | MenhirState159 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1949 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | END -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1945 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1943 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _endpos_c_, _, (c : 'tv_commands)) = _menhir_stack in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_command_core = +# 551 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( c.p_command_desc ) +# 555 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1944)) : 'freshtv1946) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1947 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1948)) : 'freshtv1950) + | MenhirState170 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1953 * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1951 * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_command), _startpos__1_), _endpos__3_, _, (_3 : 'tv_commands)) = _menhir_stack in + let _2 = () in + let _endpos = _endpos__3_ in + let _v : 'tv_commands = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 598 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkcmd ~loc:_sloc (PCsequence (_1, _3)) ) +# 579 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_commands _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1952)) : 'freshtv1954) + | MenhirState124 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1959 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 587 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1955 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 597 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState183 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState183 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run122 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState183 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState183 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState183) : 'freshtv1956) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1957 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 649 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1958)) : 'freshtv1960) + | MenhirState117 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1965 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 658 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1961 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 668 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState190 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState190 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState190 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState190 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState190) : 'freshtv1962) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1963 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 726 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1964)) : 'freshtv1966) + | MenhirState190 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1969 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 735 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1967 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 741 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, (_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 746 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__2_), _endpos__4_, _, (_4 : 'tv_commands)), _endpos__6_, _, (_6 : 'tv_commands)) = _menhir_stack in + let _5 = () in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__6_ in + let _v : 'tv_command = let _endpos = _endpos__6_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 586 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkcmd ~loc:_sloc (PClet (_2, _4, _6)) ) +# 759 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1968)) : 'freshtv1970) + | MenhirState114 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1973 * _menhir_state * Lexing.position) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1971 * _menhir_state * Lexing.position) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _menhir_s, _startpos__1_), _endpos__2_), _endpos__4_, _, (_4 : 'tv_commands)) = _menhir_stack in + let _3 = () in + let _2 = () in + let _1 = () in + let _v : 'tv_match_clause = +# 616 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( ("NIL", [], _4) ) +# 774 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_match_clause _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1972)) : 'freshtv1974) + | MenhirState201 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1977 * _menhir_state * 'tv_match_pattern)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1975 * _menhir_state * 'tv_match_pattern)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, (_1 : 'tv_match_pattern)), _endpos__3_, _, (_3 : 'tv_commands)) = _menhir_stack in + let _2 = () in + let _v : 'tv_match_clause = +# 612 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( let (cnstr, params) = _1 in + (cnstr , params, _3) ) +# 788 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_match_clause _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1976)) : 'freshtv1978) + | MenhirState491 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1981 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1979 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let ((((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_annotations)), _), _, (_4 : 'tv_method_param)), _, (_5 : 'tv_opt_type_annotation)), _endpos__7_, _, (_7 : 'tv_commands)) = _menhir_stack in + let _6 = () in + let _3 = () in + let _1 = () in + let _v : 'tv_object_field_or_method = let _endpos = _endpos__7_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 707 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( `Right ("constructor", { pMethodArguments = _4; + pMethodReturnType = _5; + pMethodKind = MKconstructor; + pMethodBody = _7; + pMethodAnnotations = _2; + pMethodLoc = (make_loc _sloc) }) ) +# 811 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_field_or_method _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1980)) : 'freshtv1982) + | MenhirState501 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1985 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 819 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1983 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 825 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation)) * Lexing.position * _menhir_state * 'tv_commands) = Obj.magic _menhir_stack in + ((let (((((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_annotations)), _, (_3 : 'tv_method_kind)), _endpos__4_, (_4 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 830 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__4_), _, (_5 : 'tv_method_param)), _, (_6 : 'tv_opt_type_annotation)), _endpos__8_, _, (_8 : 'tv_commands)) = _menhir_stack in + let _7 = () in + let _1 = () in + let _v : 'tv_object_field_or_method = let _endpos = _endpos__8_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 699 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( `Right (_4, { pMethodArguments = _5; + pMethodReturnType = _6; + pMethodKind = _3; + pMethodBody = _8; + pMethodAnnotations = _2; + pMethodLoc = (make_loc _sloc) }) ) +# 845 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_field_or_method _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1984)) : 'freshtv1986) + | _ -> + _menhir_fail () + +and _menhir_run199 : _menhir_env -> 'ttv_tail * _menhir_state * 'tv_idents -> Lexing.position -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 854 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1941 * _menhir_state * 'tv_idents) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let ((_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 864 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 868 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__2_ : Lexing.position) = _startpos in + ((let (_menhir_stack, _menhir_s, (_1 : 'tv_idents)) = _menhir_stack in + let _v : 'tv_idents = +# 316 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 :: _1 ) +# 875 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_idents _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1942) + +and _menhir_goto_object_fields_and_methods : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_object_fields_and_methods -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1939 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 886 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) * _menhir_state * 'tv_object_fields_and_methods) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | LET -> + _menhir_run474 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState503 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1937 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 898 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) * _menhir_state * 'tv_object_fields_and_methods) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState503 in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1935 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 907 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) * _menhir_state * 'tv_object_fields_and_methods) = Obj.magic _menhir_stack in + let (_endpos__9_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + ((let (((((((_menhir_stack, _endpos__1_, (_1 : 'tv_opt_logical_or_trusted), _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 914 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), (_4 : 'tv_base_layer_signature)), _endpos__6_, _, (_6 : 'tv_object_signature_expr), _startpos__6_), _startpos__7_), _, (_8 : 'tv_object_fields_and_methods)) = _menhir_stack in + let _9 = () in + let _7 = () in + let _5 = () in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__9_ in + let _v : 'tv_object_declaration = let _endpos = _endpos__9_ in + let _symbolstartpos = if _startpos__1_ != _endpos__1_ then + _startpos__1_ + else + _startpos__2_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 637 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( let fields_rev, methods_rev = _8 in + let obj_type = { pObjectBase = _4; pObjectSignature = _6; pObjectTypLoc = (make_loc _sloc) } in + _3, { + pObjectType = None; + pObjectDesc = mkobj ~loc:_sloc (POconstr { + pObjType = obj_type; + pObjKind = _1; + pObjFields = List.rev fields_rev; + pObjMethods = List.rev methods_rev; + pObjLoc = (make_loc _sloc) }); + pObjectLoc = (make_loc _sloc) } + ) +# 942 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_declaration _menhir_env _menhir_stack _endpos _v _startpos) : 'freshtv1936)) : 'freshtv1938) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState503) : 'freshtv1940) + +and _menhir_goto_idents_semi_sep : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_idents_semi_sep -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1933 * _menhir_state * 'tv_idents_semi_sep) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1931 * _menhir_state * 'tv_idents_semi_sep) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState304 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1927 * _menhir_state * 'tv_idents_semi_sep) * _menhir_state) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 973 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1925 * _menhir_state * 'tv_idents_semi_sep) * _menhir_state) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + let ((_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 983 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 987 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__3_ : Lexing.position) = _startpos in + ((let ((_menhir_stack, _menhir_s, (_1 : 'tv_idents_semi_sep)), _) = _menhir_stack in + let _2 = () in + let _v : 'tv_idents_semi_sep = +# 391 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 995 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_idents_semi_sep _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1926)) : 'freshtv1928) + | RBRACE -> + _menhir_reduce214 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1929 * _menhir_state * 'tv_idents_semi_sep) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1930)) : 'freshtv1932) + | RBRACE -> + _menhir_reduce213 _menhir_env (Obj.magic _menhir_stack) MenhirState304 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState304) : 'freshtv1934) + +and _menhir_goto_layer_slots_plus : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_layer_slots_plus -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1923 * _menhir_state * 'tv_layer_slots_plus) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1921 * _menhir_state * 'tv_layer_slots_plus) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState375 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run353 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState376 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RBRACE -> + _menhir_reduce214 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState376) : 'freshtv1922) + | RBRACE -> + _menhir_reduce213 _menhir_env (Obj.magic _menhir_stack) MenhirState375 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState375) : 'freshtv1924) + +and _menhir_goto_layer_invariant_annotation : _menhir_env -> 'ttv_tail -> Lexing.position -> 'tv_layer_invariant_annotation -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1919 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1051 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * 'tv_layer_signature_annotation)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_layer_invariant_annotation) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1917 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1059 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * 'tv_layer_signature_annotation)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__7_ : Lexing.position) = _endpos in + let ((_7 : 'tv_layer_invariant_annotation) : 'tv_layer_invariant_annotation) = _v in + ((let (((((_menhir_stack, _startpos__1_), _endpos__2_, (_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1066 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__2_), _endpos__3_, _, (_3 : 'tv_annotations)), (_4 : 'tv_layer_signature_annotation)), _endpos__6_, _, (_6 : 'tv_layer_expression), _startpos__6_) = _menhir_stack in + let _5 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__7_ in + let _v : 'tv_layer_declaration = let _endpos = _endpos__7_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 742 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2, { pLayerLoc = (make_loc _sloc); + pLayerAnnotations = _3; + pLayerInvariant = _7; + pLayerDesc = match _4 with + | None -> _6 + | Some t -> mklayer ~loc:_sloc (PLrelax (_6, t)); + } + ) +# 1085 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1915) = _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_layer_declaration) = _v in + let (_startpos : Lexing.position) = _startpos in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1913) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_layer_declaration) = _v in + let (_startpos : Lexing.position) = _startpos in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1911) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let ((_1 : 'tv_layer_declaration) : 'tv_layer_declaration) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _endpos = _endpos__1_ in + let _v : 'tv_declaration = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 228 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( fst _1, mkdecl ~loc:_sloc (PDlayer (snd _1)) ) +# 1109 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv1912)) : 'freshtv1914)) : 'freshtv1916)) : 'freshtv1918)) : 'freshtv1920) + +and _menhir_run388 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1116 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1909) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1127 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1131 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _endpos = _endpos__1_ in + let _v : 'tv_proposition = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 820 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkprop ~loc:_sloc (PPexternal _1) ) +# 1141 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_proposition _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1910) + +and _menhir_run389 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1148 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1907) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1159 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1163 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _endpos = _endpos__1_ in + let _v : 'tv_proposition = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 819 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkprop ~loc:_sloc (PPident _1) ) +# 1173 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_proposition _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1908) + +and _menhir_run385 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run382 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState385 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run352 _menhir_env (Obj.magic _menhir_stack) MenhirState385 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run351 _menhir_env (Obj.magic _menhir_stack) MenhirState385 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState385 + +and _menhir_run391 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + _menhir_run342 _menhir_env (Obj.magic _menhir_stack) MenhirState391 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState391 + +and _menhir_run393 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run382 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState393 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run352 _menhir_env (Obj.magic _menhir_stack) MenhirState393 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run351 _menhir_env (Obj.magic _menhir_stack) MenhirState393 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState393 + +and _menhir_reduce215 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _v : 'tv_opt_type_annotation = +# 728 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( None ) +# 1226 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_opt_type_annotation _menhir_env _menhir_stack _menhir_s _v + +and _menhir_run480 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState480 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState480 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState480 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState480 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState480 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState480 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState480 + +and _menhir_goto_layer_slots : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_layer_slots -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1905 * _menhir_state * Lexing.position) * _menhir_state * 'tv_layer_slots) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1901 * _menhir_state * Lexing.position) * _menhir_state * 'tv_layer_slots) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1899 * _menhir_state * Lexing.position) * _menhir_state * 'tv_layer_slots) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _, (_2 : 'tv_layer_slots)) = _menhir_stack in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_layer_expression = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 762 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayer ~loc:_sloc (PLconstr (List.rev _2)) ) +# 1280 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1900)) : 'freshtv1902) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1903 * _menhir_state * Lexing.position) * _menhir_state * 'tv_layer_slots) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1904)) : 'freshtv1906) + +and _menhir_run353 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1294 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1895 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1306 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | CLONE -> + _menhir_run357 _menhir_env (Obj.magic _menhir_stack) MenhirState354 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run356 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState354 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1893) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState354 in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1889 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | INT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1875 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1340 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1871 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1352 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_stack = (_menhir_stack, _endpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLONLESS -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1867 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1364 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | CLONE -> + _menhir_run357 _menhir_env (Obj.magic _menhir_stack) MenhirState371 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run356 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState371 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run355 _menhir_env (Obj.magic _menhir_stack) MenhirState371 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState371) : 'freshtv1868) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1869 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1386 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _, _menhir_s, _), _), _, _, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1870)) : 'freshtv1872) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1873 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1397 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _), _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1874)) : 'freshtv1876) + | UINT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1885 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1408 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1881 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1420 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_stack = (_menhir_stack, _endpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLONLESS -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1877 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1432 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | CLONE -> + _menhir_run357 _menhir_env (Obj.magic _menhir_stack) MenhirState367 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run356 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState367 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run355 _menhir_env (Obj.magic _menhir_stack) MenhirState367 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState367) : 'freshtv1878) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1879 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1454 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _, _menhir_s, _), _), _, _, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1880)) : 'freshtv1882) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1883 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 1465 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _), _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1884)) : 'freshtv1886) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1887 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1888)) : 'freshtv1890) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1891 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1892)) : 'freshtv1894) + | LPAREN -> + _menhir_run355 _menhir_env (Obj.magic _menhir_stack) MenhirState354 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState354) : 'freshtv1896) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1897 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1496 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1898) + +and _menhir_run211 : _menhir_env -> 'ttv_tail * _menhir_state * 'tv_annotations_plus -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + _menhir_run30 _menhir_env (Obj.magic _menhir_stack) MenhirState211 + | IDENT _v -> + _menhir_run20 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState211 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run14 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState211 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run13 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState211 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | STRING _v -> + _menhir_run8 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState211 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run7 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState211 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState211 + +and _menhir_goto_command : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_command -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + match _menhir_s with + | MenhirState501 | MenhirState491 | MenhirState201 | MenhirState114 | MenhirState117 | MenhirState190 | MenhirState124 | MenhirState170 | MenhirState159 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1815 * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1809 * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState170 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState170 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState170 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState170 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState170) : 'freshtv1810) + | BAR | COMMA | ELSE | END | IN | LET | RBRACE | RBRACKET | RPAREN | THEN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1811 * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _endpos_c_, _menhir_s, (c : 'tv_command), _startpos_c_) = _menhir_stack in + let _endpos = _endpos_c_ in + let _v : 'tv_commands = let _endpos = _endpos_c_ in + let _symbolstartpos = _startpos_c_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 597 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkcmd ~loc:_sloc c.p_command_desc ) +# 1596 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_commands _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1812) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1813 * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1814)) : 'freshtv1816) + | MenhirState154 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1825 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1611 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ELSE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1817 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1621 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState178 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState178 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState178 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState178 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState178) : 'freshtv1818) + | THEN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1819 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1677 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState174 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState174 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState174 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState174 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState174) : 'freshtv1820) + | BAR | COMMA | END | IN | LET | RBRACE | RBRACKET | RPAREN | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1821 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1733 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (((((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_annotations)), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1738 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _, (_4 : 'tv_eq_or_assign)), _endpos_e1_, _, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_), _endpos__9_, _, (_9 : 'tv_command), _startpos__9_) = _menhir_stack in + let _8 = () in + let _6 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__9_ in + let _v : 'tv_command_core = let _endpos = _endpos__9_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 565 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCfirst (_3, mkexp_ ~loc:_sloc e1.p_expression_desc, mkexp_ ~loc:_sloc e2.p_expression_desc, _9, None, None, _2) ) +# 1751 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1822) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1823 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1761 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1824)) : 'freshtv1826) + | MenhirState174 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((('freshtv1833 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1770 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ELSE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((('freshtv1827 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1780 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState176 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState176 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState176 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState176 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState176) : 'freshtv1828) + | BAR | COMMA | END | IN | LET | RBRACE | RBRACKET | RPAREN | SEMICOLON | THEN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((('freshtv1829 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1836 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_annotations)), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1841 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _, (_4 : 'tv_eq_or_assign)), _endpos_e1_, _, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_), _endpos__9_, _, (_9 : 'tv_command), _startpos__9_), _endpos__11_, _, (_11 : 'tv_command), _startpos__11_) = _menhir_stack in + let _10 = () in + let _8 = () in + let _6 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__11_ in + let _v : 'tv_command_core = let _endpos = _endpos__11_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 571 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCfirst (_3,mkexp_ ~loc:_sloc e1.p_expression_desc, mkexp_ ~loc:_sloc e2.p_expression_desc, _9, Some _11, None, _2) ) +# 1855 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1830) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((('freshtv1831 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1865 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1832)) : 'freshtv1834) + | MenhirState176 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((((('freshtv1837 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1874 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((((('freshtv1835 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1880 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (((((((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_annotations)), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1885 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _, (_4 : 'tv_eq_or_assign)), _endpos_e1_, _, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_), _endpos__9_, _, (_9 : 'tv_command), _startpos__9_), _endpos__11_, _, (_11 : 'tv_command), _startpos__11_), _endpos__13_, _, (_13 : 'tv_command), _startpos__13_) = _menhir_stack in + let _12 = () in + let _10 = () in + let _8 = () in + let _6 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__13_ in + let _v : 'tv_command_core = let _endpos = _endpos__13_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 574 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCfirst (_3, mkexp_ ~loc:_sloc e1.p_expression_desc, mkexp_ ~loc:_sloc e2.p_expression_desc, _9, Some _11, Some _13, _2) ) +# 1900 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1836)) : 'freshtv1838) + | MenhirState178 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((('freshtv1841 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1908 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((('freshtv1839 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1914 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_annotations)), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1919 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _, (_4 : 'tv_eq_or_assign)), _endpos_e1_, _, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_), _endpos__9_, _, (_9 : 'tv_command), _startpos__9_), _endpos__11_, _, (_11 : 'tv_command), _startpos__11_) = _menhir_stack in + let _10 = () in + let _8 = () in + let _6 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__11_ in + let _v : 'tv_command_core = let _endpos = _endpos__11_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 568 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCfirst (_3, mkexp_ ~loc:_sloc e1.p_expression_desc, mkexp_ ~loc:_sloc e2.p_expression_desc, _9, None, Some _11, _2) ) +# 1933 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1840)) : 'freshtv1842) + | MenhirState146 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((((('freshtv1845 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1941 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1945 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((((('freshtv1843 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1951 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1955 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((((((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_annotations)), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1960 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _, (_4 : 'tv_eq_or_assign)), _endpos_e1_, _, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_), _endpos__9_, (_9 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1964 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__9_), _, (_10 : 'tv_eq_or_assign)), _endpos_e3_, _, (e3 : 'tv_expression), _startpos_e3_), _endpos__13_, _, (_13 : 'tv_command), _startpos__13_) = _menhir_stack in + let _12 = () in + let _8 = () in + let _6 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__13_ in + let _v : 'tv_command_core = let _endpos = _endpos__13_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 577 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCfold (_3, mkexp_ ~loc:_sloc e1.p_expression_desc, mkexp_ ~loc:_sloc e2.p_expression_desc, _9, mkexp_ ~loc:_sloc e3.p_expression_desc, _13, _2) ) +# 1978 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1844)) : 'freshtv1846) + | MenhirState134 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1849 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1986 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1847 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1992 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (((((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_annotations)), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 1997 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _, (_4 : 'tv_eq_or_assign)), _endpos_e1_, _, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_), _endpos__9_, _, (_9 : 'tv_command), _startpos__9_) = _menhir_stack in + let _8 = () in + let _6 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__9_ in + let _v : 'tv_command_core = let _endpos = _endpos__9_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 562 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCfor (_3, mkexp_ ~loc:_sloc e1.p_expression_desc, mkexp_ ~loc:_sloc e2.p_expression_desc, _9, _2) ) +# 2010 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1848)) : 'freshtv1850) + | MenhirState120 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1857 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ELSE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1851 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState187 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState187 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState187 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState187 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState187) : 'freshtv1852) + | BAR | COMMA | END | IN | LET | RBRACE | RBRACKET | RPAREN | SEMICOLON | THEN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1853 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _menhir_s, _startpos__1_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_), _endpos__4_, _, (_4 : 'tv_command), _startpos__4_) = _menhir_stack in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__4_ in + let _v : 'tv_command_core = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 556 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCcond (mkexp_ ~loc:_sloc e.p_expression_desc, _4, None) ) +# 2085 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1854) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1855 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1856)) : 'freshtv1858) + | MenhirState187 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1861 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1859 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _menhir_s, _startpos__1_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_), _endpos__4_, _, (_4 : 'tv_command), _startpos__4_), _endpos__6_, _, (_6 : 'tv_command), _startpos__6_) = _menhir_stack in + let _5 = () in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__6_ in + let _v : 'tv_command_core = let _endpos = _endpos__6_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 557 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCcond (mkexp_ ~loc:_sloc e.p_expression_desc, _4, Some _6) ) +# 2112 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1860)) : 'freshtv1862) + | MenhirState30 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1865 * _menhir_state) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1863 * _menhir_state) * Lexing.position * _menhir_state * 'tv_command * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s), _endpos__2_, _, (_2 : 'tv_command), _startpos__2_) = _menhir_stack in + let _1 = () in + let _v : 'tv_annotation = +# 798 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PAexpr _2 ) +# 2125 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1864)) : 'freshtv1866) + | _ -> + _menhir_fail () + +and _menhir_goto_annotated_command : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_annotated_command -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v -> + match _menhir_s with + | MenhirState160 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1795 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_annotated_command) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1793 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_2 : 'tv_annotated_command) : 'tv_annotated_command) = _v in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__2_ in + let _v : 'tv_command = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 582 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkcmd ~loc:_sloc (PCassert _2) ) +# 2155 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1794)) : 'freshtv1796) + | MenhirState158 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1799 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_annotated_command) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1797 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_2 : 'tv_annotated_command) : 'tv_annotated_command) = _v in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__2_ in + let _v : 'tv_command = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 583 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkcmd ~loc:_sloc (PCdeny _2) ) +# 2179 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1798)) : 'freshtv1800) + | MenhirState183 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1803 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 2187 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands)) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_annotated_command) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1801 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 2196 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands)) = Obj.magic _menhir_stack in + let (_endpos__6_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_6 : 'tv_annotated_command) : 'tv_annotated_command) = _v in + ((let (((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, (_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 2204 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__2_), _endpos__4_, _, (_4 : 'tv_commands)) = _menhir_stack in + let _5 = () in + let _3 = () in + let _1 = () in + let _endpos = _endpos__6_ in + let _v : 'tv_annotated_command = let _endpos = _endpos__6_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 590 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkcmd ~loc:_sloc (PClet (_2, _4, _6)) ) +# 2216 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotated_command _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1802)) : 'freshtv1804) + | MenhirState121 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1807 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_annotated_command) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1805 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_2 : 'tv_annotated_command) : 'tv_annotated_command) = _v in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__2_ in + let _v : 'tv_command = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 584 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkcmd ~loc:_sloc (PCghost _2) ) +# 2240 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1806)) : 'freshtv1808) + | _ -> + _menhir_fail () + +and _menhir_goto_opt_semi : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_opt_semi -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + match _menhir_s with + | MenhirState95 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1733 * _menhir_state * Lexing.position) * _menhir_state * 'tv_struct_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1729 * _menhir_state * Lexing.position) * _menhir_state * 'tv_struct_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1727 * _menhir_state * Lexing.position) * _menhir_state * 'tv_struct_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _menhir_s, _startpos__1_), _, (_2 : 'tv_struct_fields)), _, (_3 : 'tv_opt_semi)) = _menhir_stack in + let _4 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__4_ in + let _v : 'tv_expression = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 533 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEstruct (List.rev _2)) ) +# 2275 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1728)) : 'freshtv1730) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1731 * _menhir_state * Lexing.position) * _menhir_state * 'tv_struct_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1732)) : 'freshtv1734) + | MenhirState247 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1741 * _menhir_state * Lexing.position) * _menhir_state * 'tv_field_declarations) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1737 * _menhir_state * Lexing.position) * _menhir_state * 'tv_field_declarations) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1735 * _menhir_state * Lexing.position) * _menhir_state * 'tv_field_declarations) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _menhir_s, _startpos__1_), _, (_2 : 'tv_field_declarations)), _, (_3 : 'tv_opt_semi)) = _menhir_stack in + let _4 = () in + let _1 = () in + let _endpos = _endpos__4_ in + let _v : 'tv_type_definition = +# 279 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PTsingleton (List.rev _2) ) +# 2306 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_definition _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1736)) : 'freshtv1738) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1739 * _menhir_state * Lexing.position) * _menhir_state * 'tv_field_declarations) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1740)) : 'freshtv1742) + | MenhirState286 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1749 * _menhir_state * Lexing.position) * _menhir_state * 'tv_object_signature_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1745 * _menhir_state * Lexing.position) * _menhir_state * 'tv_object_signature_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1743 * _menhir_state * Lexing.position) * _menhir_state * 'tv_object_signature_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _menhir_s, _startpos__1_), _, (_2 : 'tv_object_signature_fields)), _, (_3 : 'tv_opt_semi)) = _menhir_stack in + let _4 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__4_ in + let _v : 'tv_object_signature_expr = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 360 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mksig ~loc:_sloc (PSconstr (List.rev _2)) ) +# 2341 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_expr _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1744)) : 'freshtv1746) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1747 * _menhir_state * Lexing.position) * _menhir_state * 'tv_object_signature_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1748)) : 'freshtv1750) + | MenhirState304 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1779 * _menhir_state * 'tv_idents_semi_sep) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1777 * _menhir_state * 'tv_idents_semi_sep) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, (_1 : 'tv_idents_semi_sep)), _, (_2 : 'tv_opt_semi)) = _menhir_stack in + let _v : 'tv_idents_semi = +# 387 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( List.rev _1 ) +# 2360 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1775) = _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_idents_semi) = _v in + ((let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + match _menhir_s with + | MenhirState302 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1757 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1753 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1751 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + let (_endpos__6_ : Lexing.position) = _endpos in + ((let ((((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_object_signature_expr), _startpos__1_), _endpos__3_, _startpos__3_), _startpos__4_), _, (_5 : 'tv_idents_semi)) = _menhir_stack in + let _6 = () in + let _4 = () in + let _3 = () in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__6_ in + let _v : 'tv_object_signature_expr = let _endpos = _endpos__6_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 364 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mksig ~loc:_sloc (PSlogicize (_1, _5)) ) +# 2395 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_expr _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1752)) : 'freshtv1754) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1755 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1756)) : 'freshtv1758) + | MenhirState312 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1765 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1761 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1759 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + let (_endpos__6_ : Lexing.position) = _endpos in + ((let ((((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_object_signature_expr), _startpos__1_), _startpos__3_), _startpos__4_), _, (_5 : 'tv_idents_semi)) = _menhir_stack in + let _6 = () in + let _4 = () in + let _3 = () in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__6_ in + let _v : 'tv_object_signature_expr = let _endpos = _endpos__6_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 362 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mksig ~loc:_sloc (PSghostize (_1, _5)) ) +# 2432 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_expr _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1760)) : 'freshtv1762) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1763 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1764)) : 'freshtv1766) + | MenhirState317 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1773 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1769 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1767 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + let (_endpos__5_ : Lexing.position) = _endpos in + ((let ((((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_object_signature_expr), _startpos__1_), _startpos__2_), _startpos__3_), _, (_4 : 'tv_idents_semi)) = _menhir_stack in + let _5 = () in + let _3 = () in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__5_ in + let _v : 'tv_object_signature_expr = let _endpos = _endpos__5_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 366 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mksig ~loc:_sloc (PSminus (_1, _4)) ) +# 2468 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_expr _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1768)) : 'freshtv1770) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1771 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) * Lexing.position) * _menhir_state * 'tv_idents_semi) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1772)) : 'freshtv1774) + | _ -> + _menhir_fail ()) : 'freshtv1776)) : 'freshtv1778)) : 'freshtv1780) + | MenhirState331 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1787 * _menhir_state * Lexing.position) * _menhir_state * 'tv_layer_signature_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1783 * _menhir_state * Lexing.position) * _menhir_state * 'tv_layer_signature_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1781 * _menhir_state * Lexing.position) * _menhir_state * 'tv_layer_signature_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _menhir_s, _startpos__1_), _, (_2 : 'tv_layer_signature_fields)), _, (_3 : 'tv_opt_semi)) = _menhir_stack in + let _4 = () in + let _1 = () in + let _endpos = _endpos__4_ in + let _v : 'tv_layer_signature = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 402 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayersign ~loc:_sloc (PLSconstr (List.rev _2)) ) +# 2504 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_signature _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1782)) : 'freshtv1784) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1785 * _menhir_state * Lexing.position) * _menhir_state * 'tv_layer_signature_fields) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1786)) : 'freshtv1788) + | MenhirState375 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1791 * _menhir_state * 'tv_layer_slots_plus) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1789 * _menhir_state * 'tv_layer_slots_plus) * _menhir_state * 'tv_opt_semi) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, (_1 : 'tv_layer_slots_plus)), _, (_2 : 'tv_opt_semi)) = _menhir_stack in + let _v : 'tv_layer_slots = +# 772 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _1 ) +# 2523 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_slots _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1790)) : 'freshtv1792) + | _ -> + _menhir_fail () + +and _menhir_goto_event_declarations : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_event_declarations -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1725) * _menhir_state * 'tv_opt_bar) * Lexing.position * _menhir_state * 'tv_event_declarations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BAR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1713 * Lexing.position * _menhir_state * 'tv_event_declarations) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run428 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState439 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState439) : 'freshtv1714) + | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1721) * _menhir_state * 'tv_opt_bar) * Lexing.position * _menhir_state * 'tv_event_declarations) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, (_2 : 'tv_opt_bar)), _endpos__3_, _, (_3 : 'tv_event_declarations)) = _menhir_stack in + let _1 = () in + let _endpos = _endpos__3_ in + let _v : 'tv_events_declaration = +# 320 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 ) +# 2558 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1719) = _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_events_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1717 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_events_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1715 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let ((_2 : 'tv_events_declaration) : 'tv_events_declaration) = _v in + ((let (_menhir_stack, _endpos__1_, (_1 : 'tv_declarations)) = _menhir_stack in + let _endpos = _endpos__2_ in + let _v : 'tv_declarations = +# 219 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 @ _1 ) +# 2577 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_declarations _menhir_env _menhir_stack _endpos _v) : 'freshtv1716)) : 'freshtv1718)) : 'freshtv1720)) : 'freshtv1722) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1723) * _menhir_state * 'tv_opt_bar) * Lexing.position * _menhir_state * 'tv_event_declarations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1724)) : 'freshtv1726) + +and _menhir_goto_idents : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_idents -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + match _menhir_s with + | MenhirState193 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1705 * _menhir_state * 'tv_idents) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run199 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DOUBLEARROW -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1701 * _menhir_state * 'tv_idents) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, (_1 : 'tv_idents)) = _menhir_stack in + let _v : 'tv_match_pattern_tail = +# 629 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PPTother (List.rev _1) ) +# 2607 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_match_pattern_tail _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1702) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1703 * _menhir_state * 'tv_idents) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1704)) : 'freshtv1706) + | MenhirState259 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1711 * Lexing.position) * _menhir_state * 'tv_idents) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1707 * Lexing.position) * _menhir_state * 'tv_idents) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState261 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState261 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState261 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState261 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState261 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState261 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState261) : 'freshtv1708) + | IDENT _v -> + _menhir_run199 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1709 * Lexing.position) * _menhir_state * 'tv_idents) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1710)) : 'freshtv1712) + | _ -> + _menhir_fail () + +and _menhir_goto_object_field_or_method : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_object_field_or_method -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + match _menhir_s with + | MenhirState503 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1695 * _menhir_state * 'tv_object_fields_and_methods) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_object_field_or_method) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1693 * _menhir_state * 'tv_object_fields_and_methods) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_2 : 'tv_object_field_or_method) : 'tv_object_field_or_method) = _v in + ((let (_menhir_stack, _menhir_s, (_1 : 'tv_object_fields_and_methods)) = _menhir_stack in + let _v : 'tv_object_fields_and_methods = +# 678 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( cons_either_double_list _2 _1 ) +# 2673 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_fields_and_methods _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1694)) : 'freshtv1696) + | MenhirState473 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1699) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_object_field_or_method) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1697) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_object_field_or_method) : 'tv_object_field_or_method) = _v in + ((let _v : 'tv_object_fields_and_methods = +# 676 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( either_to_double_list _1 ) +# 2688 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_fields_and_methods _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1698)) : 'freshtv1700) + | _ -> + _menhir_fail () + +and _menhir_goto_comma_sep_expressions : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_comma_sep_expressions -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v) in + match _menhir_s with + | MenhirState102 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1683 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_comma_sep_expressions) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1681 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_comma_sep_expressions) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_expression), _startpos__1_), _endpos__3_, _, (_3 : 'tv_comma_sep_expressions)) = _menhir_stack in + let _2 = () in + let _endpos = _endpos__3_ in + let _v : 'tv_comma_sep_expressions = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 538 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEpair (_1, _3)) ) +# 2712 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_comma_sep_expressions _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1682)) : 'freshtv1684) + | MenhirState33 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1691 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_comma_sep_expressions) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1687 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_comma_sep_expressions) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1685 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_comma_sep_expressions) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_comma_sep_expressions)) = _menhir_stack in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_atom = +# 498 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 ) +# 2737 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_atom _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1686)) : 'freshtv1688) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1689 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_comma_sep_expressions) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1690)) : 'freshtv1692) + | _ -> + _menhir_fail () + +and _menhir_goto_struct_fields : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_struct_fields -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1679 * _menhir_state * Lexing.position) * _menhir_state * 'tv_struct_fields) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1677 * _menhir_state * 'tv_struct_fields) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState95 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run36 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState96 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RBRACE -> + _menhir_reduce214 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState96) : 'freshtv1678) + | RBRACE -> + _menhir_reduce213 _menhir_env (Obj.magic _menhir_stack) MenhirState95 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState95) : 'freshtv1680) + +and _menhir_run54 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState54 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState54 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState54 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState54 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState54 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState54 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState54 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState54 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState54 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState54 + +and _menhir_run72 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState72 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState72 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState72 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState72 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState72 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState72 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState72 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState72 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState72 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState72 + +and _menhir_run56 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState56 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState56 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState56 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState56 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState56 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState56 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState56 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState56 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState56 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState56 + +and _menhir_run58 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState58 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState58 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState58 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState58 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState58 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState58 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState58 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState58 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState58 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState58 + +and _menhir_run60 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState60 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState60 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState60 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState60 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState60 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState60 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState60 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState60 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState60 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState60 + +and _menhir_run68 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState68 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState68 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState68 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState68 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState68 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState68 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState68 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState68 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState68 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState68 + +and _menhir_run62 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState62 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState62 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState62 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState62 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState62 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState62 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState62 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState62 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState62 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState62 + +and _menhir_run64 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState64 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState64 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState64 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState64 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState64 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState64 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState64 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState64 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState64 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState64 + +and _menhir_run66 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _startpos -> + let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState66 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState66 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState66 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState66 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState66 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState66 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState66 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState66 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState66 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState66 + +and _menhir_run77 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState77 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState77 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState77 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState77 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState77 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState77 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState77 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState77 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState77 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState77 + +and _menhir_run79 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState79 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState79 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState79 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState79 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState79 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState79 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState79 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState79 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState79 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState79 + +and _menhir_run81 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState81 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState81 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState81 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState81 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState81 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState81 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState81 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState81 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState81 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState81 + +and _menhir_run83 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState83 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState83 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState83 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState83 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState83 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState83 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState83 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState83 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState83 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState83 + +and _menhir_run85 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState85 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState85 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState85 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState85 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState85 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState85 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState85 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState85 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState85 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState85 + +and _menhir_run87 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState87 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState87 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState87 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState87 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState87 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState87 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState87 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState87 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState87 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState87 + +and _menhir_run89 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState89 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState89 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState89 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState89 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState89 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState89 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState89 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState89 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState89 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState89 + +and _menhir_run70 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState70 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState70 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState70 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState70 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState70 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState70 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState70 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState70 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState70 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState70 + +and _menhir_run74 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState74 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState74 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState74 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState74 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState74 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState74 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState74 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState74 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState74 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState74 + +and _menhir_run303 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3289 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1675) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3300 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3304 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _v : 'tv_idents_semi_sep = +# 390 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 3310 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_idents_semi_sep _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1676) + +and _menhir_goto_object_declaration : _menhir_env -> 'ttv_tail -> Lexing.position -> 'tv_object_declaration -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v _startpos -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1673) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_object_declaration) = _v in + let (_startpos : Lexing.position) = _startpos in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1671) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let ((_1 : 'tv_object_declaration) : 'tv_object_declaration) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _endpos = _endpos__1_ in + let _v : 'tv_declaration = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 227 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( fst _1, mkdecl ~loc:_sloc (PDobject (snd _1)) ) +# 3333 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv1672)) : 'freshtv1674) + +and _menhir_goto_layer_obj_inst : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_layer_obj_inst -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1669 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3343 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_layer_obj_inst) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1667 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3351 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_layer_obj_inst) : 'tv_layer_obj_inst) = _v in + ((let (_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3358 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_) = _menhir_stack in + let _2 = () in + let _v : 'tv_layer_slot = +# 779 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _1, _3 ) +# 3364 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1665) = _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_layer_slot) = _v in + ((match _menhir_s with + | MenhirState376 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1659 * _menhir_state * 'tv_layer_slots_plus) * _menhir_state) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_layer_slot) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1657 * _menhir_state * 'tv_layer_slots_plus) * _menhir_state) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_layer_slot) : 'tv_layer_slot) = _v in + ((let ((_menhir_stack, _menhir_s, (_1 : 'tv_layer_slots_plus)), _) = _menhir_stack in + let _2 = () in + let _v : 'tv_layer_slots_plus = +# 776 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 3385 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_slots_plus _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1658)) : 'freshtv1660) + | MenhirState352 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1663) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_layer_slot) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1661) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_layer_slot) : 'tv_layer_slot) = _v in + ((let _v : 'tv_layer_slots_plus = +# 775 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 3400 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_slots_plus _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1662)) : 'freshtv1664) + | _ -> + _menhir_fail ()) : 'freshtv1666)) : 'freshtv1668)) : 'freshtv1670) + +and _menhir_run361 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run337 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState361 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run326 _menhir_env (Obj.magic _menhir_stack) MenhirState361 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState361 + +and _menhir_goto_layer_expression : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_layer_expression -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + match _menhir_s with + | MenhirState351 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1635 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | AT -> + _menhir_run393 _menhir_env (Obj.magic _menhir_stack) + | COLON -> + _menhir_run391 _menhir_env (Obj.magic _menhir_stack) + | COLONGREATER -> + _menhir_run385 _menhir_env (Obj.magic _menhir_stack) + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1631 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1629 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_layer_expression), _startpos__2_) = _menhir_stack in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_layer_expression = +# 769 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 ) +# 3452 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1630)) : 'freshtv1632) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1633 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1634)) : 'freshtv1636) + | MenhirState385 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1641 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | AT -> + _menhir_run393 _menhir_env (Obj.magic _menhir_stack) + | COLON -> + _menhir_run391 _menhir_env (Obj.magic _menhir_stack) + | COLONGREATER -> + _menhir_run385 _menhir_env (Obj.magic _menhir_stack) + | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1637 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run389 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState387 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | STRING _v -> + _menhir_run388 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState387 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState387) : 'freshtv1638) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1639 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1640)) : 'freshtv1642) + | MenhirState393 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1647 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | AT -> + _menhir_run393 _menhir_env (Obj.magic _menhir_stack) + | COLON -> + _menhir_run391 _menhir_env (Obj.magic _menhir_stack) + | COLONGREATER -> + _menhir_run385 _menhir_env (Obj.magic _menhir_stack) + | ASSERT | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | RPAREN | SIGNATURE | TRUSTED | TYPE | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1643 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_layer_expression), _startpos__1_), _endpos__3_, _, (_3 : 'tv_layer_expression), _startpos__3_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_layer_expression = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 766 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayer ~loc:_sloc (PLinst (_1, _3)) ) +# 3520 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1644) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1645 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1646)) : 'freshtv1648) + | MenhirState350 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1655 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3535 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * 'tv_layer_signature_annotation)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1649) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run389 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState396 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | STRING _v -> + _menhir_run388 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState396 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState396) : 'freshtv1650) + | AT -> + _menhir_run393 _menhir_env (Obj.magic _menhir_stack) + | COLON -> + _menhir_run391 _menhir_env (Obj.magic _menhir_stack) + | COLONGREATER -> + _menhir_run385 _menhir_env (Obj.magic _menhir_stack) + | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1651) = Obj.magic _menhir_stack in + ((let (_, _endpos) = Obj.magic _menhir_stack in + let _v : 'tv_layer_invariant_annotation = +# 756 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( None ) +# 3569 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_invariant_annotation _menhir_env _menhir_stack _endpos _v) : 'freshtv1652) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1653 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3579 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * 'tv_layer_signature_annotation)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1654)) : 'freshtv1656) + | _ -> + _menhir_fail () + +and _menhir_run479 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3589 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + _menhir_run480 _menhir_env (Obj.magic _menhir_stack) MenhirState479 + | COMMA | RPAREN -> + _menhir_reduce215 _menhir_env (Obj.magic _menhir_stack) MenhirState479 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState479 + +and _menhir_goto_method_param : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_method_param -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + match _menhir_s with + | MenhirState476 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1625 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state) * _menhir_state * 'tv_method_param) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + _menhir_run480 _menhir_env (Obj.magic _menhir_stack) MenhirState489 + | EQUAL -> + _menhir_reduce215 _menhir_env (Obj.magic _menhir_stack) MenhirState489 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState489) : 'freshtv1626) + | MenhirState494 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1627 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3628 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_method_param) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + _menhir_run480 _menhir_env (Obj.magic _menhir_stack) MenhirState499 + | EQUAL -> + _menhir_reduce215 _menhir_env (Obj.magic _menhir_stack) MenhirState499 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState499) : 'freshtv1628) + | _ -> + _menhir_fail () + +and _menhir_run351 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run382 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState351 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run352 _menhir_env (Obj.magic _menhir_stack) MenhirState351 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run351 _menhir_env (Obj.magic _menhir_stack) MenhirState351 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState351 + +and _menhir_run352 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run353 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState352 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1623) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState352 in + ((let _v : 'tv_layer_slots = +# 771 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [] ) +# 3676 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_slots _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1624) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState352 + +and _menhir_run382 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3687 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1621) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3698 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3702 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_layer_expression = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 760 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayer ~loc:_sloc (PLname _1) ) +# 3713 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1622) + +and _menhir_goto_type_definition : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_type_definition -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1619 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3723 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_type_definition) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1617 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3732 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) = Obj.magic _menhir_stack in + let (_endpos__5_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_5 : 'tv_type_definition) : 'tv_type_definition) = _v in + ((let (((_menhir_stack, _startpos__1_), _endpos__2_, (_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3740 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__2_), _endpos__3_, _, (_3 : 'tv_annotations)) = _menhir_stack in + let _4 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__5_ in + let _v : 'tv_type_declaration = let _endpos = _endpos__5_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 255 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2, mkfotyp ~loc:_sloc (PTdata (_5, _3)) ) +# 3752 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_declaration _menhir_env _menhir_stack _endpos _v _startpos) : 'freshtv1618)) : 'freshtv1620) + +and _menhir_goto_annotations_plus : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_annotations_plus -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + match _menhir_s with + | MenhirState11 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1603 * _menhir_state * Lexing.position) * _menhir_state * 'tv_annotations_plus) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COMMA -> + _menhir_run211 _menhir_env (Obj.magic _menhir_stack) + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1599 * _menhir_state * Lexing.position) * _menhir_state * 'tv_annotations_plus) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1597 * _menhir_state * Lexing.position) * _menhir_state * 'tv_annotations_plus) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _, (_2 : 'tv_annotations_plus)) = _menhir_stack in + let _3 = () in + let _1 = () in + let _v : 'tv_annotation_arguments = +# 815 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( List.rev _2 ) +# 3782 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation_arguments _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1598)) : 'freshtv1600) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1601 * _menhir_state * Lexing.position) * _menhir_state * 'tv_annotations_plus) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1602)) : 'freshtv1604) + | MenhirState6 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1615 * _menhir_state * Lexing.position) * Lexing.position) * _menhir_state * 'tv_annotations_plus) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COMMA -> + _menhir_run211 _menhir_env (Obj.magic _menhir_stack) + | RBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1611 * _menhir_state * Lexing.position) * Lexing.position) * _menhir_state * 'tv_annotations_plus) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_stack = (_menhir_stack, _endpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1607 * _menhir_state * Lexing.position) * Lexing.position) * _menhir_state * 'tv_annotations_plus) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1605 * _menhir_state * Lexing.position) * Lexing.position) * _menhir_state * 'tv_annotations_plus) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__5_ : Lexing.position) = _endpos in + ((let ((((_menhir_stack, _menhir_s, _startpos__1_), _startpos__2_), _, (_3 : 'tv_annotations_plus)), _endpos__4_) = _menhir_stack in + let _5 = () in + let _4 = () in + let _2 = () in + let _1 = () in + let _endpos = _endpos__5_ in + let _v : 'tv_annotations = +# 790 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( List.rev _3 ) +# 3825 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotations _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1606)) : 'freshtv1608) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1609 * _menhir_state * Lexing.position) * Lexing.position) * _menhir_state * 'tv_annotations_plus) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1610)) : 'freshtv1612) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1613 * _menhir_state * Lexing.position) * Lexing.position) * _menhir_state * 'tv_annotations_plus) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1614)) : 'freshtv1616) + | _ -> + _menhir_fail () + +and _menhir_goto_annotation_arguments : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_annotation_arguments -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + match _menhir_s with + | MenhirState20 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1591 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3853 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_annotation_arguments) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1589 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3861 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_2 : 'tv_annotation_arguments) : 'tv_annotation_arguments) = _v in + ((let (_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3868 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_) = _menhir_stack in + let _v : 'tv_annotation = +# 803 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PAclause (_1, _2) ) +# 3873 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1590)) : 'freshtv1592) + | MenhirState8 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1595 * Lexing.position * _menhir_state * ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 3881 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_annotation_arguments) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1593 * Lexing.position * _menhir_state * ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 3889 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_2 : 'tv_annotation_arguments) : 'tv_annotation_arguments) = _v in + ((let (_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 3896 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_) = _menhir_stack in + let _v : 'tv_annotation = +# 804 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PAclause (_1, _2) ) +# 3901 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1594)) : 'freshtv1596) + | _ -> + _menhir_fail () + +and _menhir_goto_command_core : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_command_core -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + match _menhir_s with + | MenhirState121 | MenhirState183 | MenhirState158 | MenhirState160 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1583) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_command_core) = _v in + let (_startpos : Lexing.position) = _startpos in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1581) = Obj.magic _menhir_stack in + let (_endpos_c_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((c : 'tv_command_core) : 'tv_command_core) = _v in + let (_startpos_c_ : Lexing.position) = _startpos in + ((let _endpos = _endpos_c_ in + let _v : 'tv_annotated_command = let _endpos = _endpos_c_ in + let _symbolstartpos = _startpos_c_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 589 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkcmd ~loc:_sloc c) +# 3930 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotated_command _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1582)) : 'freshtv1584) + | MenhirState501 | MenhirState491 | MenhirState30 | MenhirState201 | MenhirState114 | MenhirState117 | MenhirState190 | MenhirState187 | MenhirState120 | MenhirState124 | MenhirState134 | MenhirState146 | MenhirState178 | MenhirState176 | MenhirState174 | MenhirState154 | MenhirState170 | MenhirState159 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1587) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_command_core) = _v in + let (_startpos : Lexing.position) = _startpos in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1585) = Obj.magic _menhir_stack in + let (_endpos_c_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((c : 'tv_command_core) : 'tv_command_core) = _v in + let (_startpos_c_ : Lexing.position) = _startpos in + ((let _startpos = _startpos_c_ in + let _endpos = _endpos_c_ in + let _v : 'tv_command = let _endpos = _endpos_c_ in + let _symbolstartpos = _startpos_c_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 580 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkcmd ~loc:_sloc c ) +# 3954 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1586)) : 'freshtv1588) + | _ -> + _menhir_fail () + +and _menhir_run122 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1577 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3973 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1573 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 3985 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState124 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState124 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState124 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState124 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState124) : 'freshtv1574) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1575 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4043 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1576)) : 'freshtv1578) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1579 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1580) + +and _menhir_goto_method_parameters : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_method_parameters -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1571 * _menhir_state * Lexing.position) * _menhir_state * 'tv_method_parameters) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COMMA -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1563 * _menhir_state * 'tv_method_parameters) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run479 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState485 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState485) : 'freshtv1564) + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1567 * _menhir_state * Lexing.position) * _menhir_state * 'tv_method_parameters) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1565 * _menhir_state * Lexing.position) * _menhir_state * 'tv_method_parameters) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _, (_2 : 'tv_method_parameters)) = _menhir_stack in + let _3 = () in + let _1 = () in + let _v : 'tv_method_param = +# 718 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( List.rev _2 ) +# 4089 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_param _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1566)) : 'freshtv1568) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1569 * _menhir_state * Lexing.position) * _menhir_state * 'tv_method_parameters) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1570)) : 'freshtv1572) + +and _menhir_goto_eq_or_assign : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_eq_or_assign -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + match _menhir_s with + | MenhirState127 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1517 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4109 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState130 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState130 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState130 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState130 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState130 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState130 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState130 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState130 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState130 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState130) : 'freshtv1518) + | MenhirState137 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1519 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4141 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState138 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState138 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState138 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState138 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState138 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState138 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState138 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState138 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState138 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState138) : 'freshtv1520) + | MenhirState143 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((((('freshtv1521 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4173 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4177 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState144 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState144 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState144 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState144 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState144 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState144 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState144 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState144 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState144 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState144) : 'freshtv1522) + | MenhirState149 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1523 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4209 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState150 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState150 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState150 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState150 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState150 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState150 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState150 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState150 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState150 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState150) : 'freshtv1524) + | MenhirState402 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1537 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4241 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STRING _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1533 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4251 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4257 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | STRING _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1527 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4269 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4273 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4279 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1525 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4287 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4291 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__7_ : Lexing.position) = _endpos in + let ((_7 : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4297 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4301 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__7_ : Lexing.position) = _startpos in + ((let ((((((_menhir_stack, _endpos__1_, _startpos__1_), _, _startpos__2_), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4307 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _endpos__4_, _, (_4 : 'tv_annotations)), _, (_5 : 'tv_eq_or_assign)), _endpos__6_, (_6 : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4311 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__6_) = _menhir_stack in + let _2 = () in + let _1 = () in + let _endpos = _endpos__7_ in + let _v : 'tv_external_declaration = let _endpos = _endpos__7_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 240 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3, mkdecl ~loc:_sloc (PDexternal_type (_6, Some _7, _4)) ) +# 4322 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_external_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv1526)) : 'freshtv1528) + | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1529 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4330 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4334 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((((_menhir_stack, _endpos__1_, _startpos__1_), _, _startpos__2_), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4339 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _endpos__4_, _, (_4 : 'tv_annotations)), _, (_5 : 'tv_eq_or_assign)), _endpos__6_, (_6 : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4343 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__6_) = _menhir_stack in + let _2 = () in + let _1 = () in + let _endpos = _endpos__6_ in + let _v : 'tv_external_declaration = let _endpos = _endpos__6_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 238 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3, mkdecl ~loc:_sloc (PDexternal_type (_6, None, _4)) ) +# 4354 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_external_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv1530) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1531 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4364 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4368 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1532)) : 'freshtv1534) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1535 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4379 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1536)) : 'freshtv1538) + | MenhirState412 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1545 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4388 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STRING _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1541 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4398 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4404 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1539 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4412 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + let (_endpos__10_ : Lexing.position) = _endpos in + let ((_10 : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4418 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4422 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__10_ : Lexing.position) = _startpos in + ((let ((((((((_menhir_stack, _endpos__1_, _startpos__1_), _endpos__2_, _, _startpos__2_), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4428 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _endpos__4_, _, (_4 : 'tv_annotations)), _endpos__6_, _, (_6 : 'tv_type_expression), _startpos__6_), _), _endpos__8_, _, (_8 : 'tv_type_expression), _startpos__8_), _, (_9 : 'tv_eq_or_assign)) = _menhir_stack in + let _7 = () in + let _5 = () in + let _2 = () in + let _1 = () in + let _endpos = _endpos__10_ in + let _v : 'tv_external_declaration = let _endpos = _endpos__10_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 245 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3, mkdecl ~loc:_sloc (PDexternal_function (_10, _6, _8, _4)) ) +# 4441 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_external_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv1540)) : 'freshtv1542) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv1543 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4451 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1544)) : 'freshtv1546) + | MenhirState410 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1553 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4460 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STRING _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1549 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4470 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4476 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1547 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4484 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + let (_endpos__8_ : Lexing.position) = _endpos in + let ((_8 : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4490 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4494 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__8_ : Lexing.position) = _startpos in + ((let ((((((_menhir_stack, _endpos__1_, _startpos__1_), _endpos__2_, _, _startpos__2_), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4500 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _endpos__4_, _, (_4 : 'tv_annotations)), _endpos__6_, _, (_6 : 'tv_type_expression), _startpos__6_), _, (_7 : 'tv_eq_or_assign)) = _menhir_stack in + let _5 = () in + let _2 = () in + let _1 = () in + let _endpos = _endpos__8_ in + let _v : 'tv_external_declaration = let _endpos = _endpos__8_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 242 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3, mkdecl ~loc:_sloc (PDexternal_const (_8, _6, _4)) ) +# 4512 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_external_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv1548)) : 'freshtv1550) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1551 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4522 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1552)) : 'freshtv1554) + | MenhirState420 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1561 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4531 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STRING _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1557 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4541 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4547 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1555 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4555 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + let (_endpos__7_ : Lexing.position) = _endpos in + let ((_7 : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4561 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 4565 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__7_ : Lexing.position) = _startpos in + ((let (((((_menhir_stack, _endpos__1_, _startpos__1_), _, _startpos__2_), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4571 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _endpos__5_, _, (_5 : 'tv_type_expression), _startpos__5_), _, (_6 : 'tv_eq_or_assign)) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _endpos = _endpos__7_ in + let _v : 'tv_external_declaration = let _endpos = _endpos__7_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 247 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3, mkdecl ~loc:_sloc (PDexternal_prop (_7, _5)) ) +# 4583 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_external_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv1556)) : 'freshtv1558) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1559 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4593 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1560)) : 'freshtv1562) + | _ -> + _menhir_fail () + +and _menhir_goto_object_signature_fields : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_object_signature_fields -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1515 * _menhir_state * Lexing.position) * _menhir_state * 'tv_object_signature_fields) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1513 * _menhir_state * 'tv_object_signature_fields) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState286 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | CONST -> + _menhir_run284 _menhir_env (Obj.magic _menhir_stack) MenhirState287 + | CONSTRUCTOR -> + _menhir_run279 _menhir_env (Obj.magic _menhir_stack) MenhirState287 + | GHOST -> + _menhir_run277 _menhir_env (Obj.magic _menhir_stack) MenhirState287 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LOGICAL -> + _menhir_run276 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState287 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | REFINED -> + _menhir_run274 _menhir_env (Obj.magic _menhir_stack) MenhirState287 + | RBRACE -> + _menhir_reduce214 _menhir_env (Obj.magic _menhir_stack) + | IDENT _ -> + _menhir_reduce166 _menhir_env (Obj.magic _menhir_stack) MenhirState287 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState287) : 'freshtv1514) + | RBRACE -> + _menhir_reduce213 _menhir_env (Obj.magic _menhir_stack) MenhirState286 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState286) : 'freshtv1516) + +and _menhir_reduce213 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _v : 'tv_opt_semi = +# 829 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( () ) +# 4646 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_opt_semi _menhir_env _menhir_stack _menhir_s _v + +and _menhir_reduce214 : _menhir_env -> 'ttv_tail * _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let (_menhir_stack, _menhir_s) = _menhir_stack in + let _1 = () in + let _v : 'tv_opt_semi = +# 830 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( () ) +# 4657 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_opt_semi _menhir_env _menhir_stack _menhir_s _v + +and _menhir_goto_event_parameters : _menhir_env -> 'ttv_tail -> Lexing.position -> 'tv_event_parameters -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v -> + let _menhir_stack = (_menhir_stack, _endpos, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1511 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4668 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * 'tv_event_parameters) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | LPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1495) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1491 * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4688 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1487 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4700 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState432 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState432 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState432 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState432 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState432 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState432 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState432) : 'freshtv1488) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1489 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4728 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv1490)) : 'freshtv1492) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1493 * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv1494)) : 'freshtv1496) + | BAR | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1507 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4742 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * 'tv_event_parameters) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4747 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_), _endpos__2_, (_2 : 'tv_event_parameters)) = _menhir_stack in + let _endpos = _endpos__2_ in + let _v : 'tv_event_declaration = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 330 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( (_1, mkdecl ~loc:_sloc (PDevent (List.rev _2))) ) +# 4756 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1505) = _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_event_declaration) = _v in + ((match _menhir_s with + | MenhirState439 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1499 * Lexing.position * _menhir_state * 'tv_event_declarations)) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_event_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1497 * Lexing.position * _menhir_state * 'tv_event_declarations)) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_event_declaration) : 'tv_event_declaration) = _v in + ((let (_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_event_declarations)) = _menhir_stack in + let _2 = () in + let _endpos = _endpos__3_ in + let _v : 'tv_event_declarations = +# 325 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 4781 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_event_declarations _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1498)) : 'freshtv1500) + | MenhirState427 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1503) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_event_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1501) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_event_declaration) : 'tv_event_declaration) = _v in + ((let _endpos = _endpos__1_ in + let _v : 'tv_event_declarations = +# 324 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 4799 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_event_declarations _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1502)) : 'freshtv1504) + | _ -> + _menhir_fail ()) : 'freshtv1506)) : 'freshtv1508) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1509 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4811 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * 'tv_event_parameters) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _, _), _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1510)) : 'freshtv1512) + +and _menhir_run194 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4819 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1485) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4830 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4834 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _v : 'tv_idents = +# 315 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 4840 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_idents _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1486) + +and _menhir_goto_match_pattern_tail : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_match_pattern_tail -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1483 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4850 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_match_pattern_tail) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1481 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4858 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_2 : 'tv_match_pattern_tail) : 'tv_match_pattern_tail) = _v in + ((let (_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 4865 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_) = _menhir_stack in + let _v : 'tv_match_pattern = +# 621 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( match _2 with + | PPTcons x -> ("CONS", [x; _1]) + | PPTother xs -> (_1, xs) ) +# 4872 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1479) = _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_match_pattern) = _v in + ((let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1477 * _menhir_state * 'tv_match_pattern) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | DOUBLEARROW -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1473 * _menhir_state * 'tv_match_pattern) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState201 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState201 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState201 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState201 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState201) : 'freshtv1474) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1475 * _menhir_state * 'tv_match_pattern) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1476)) : 'freshtv1478)) : 'freshtv1480)) : 'freshtv1482)) : 'freshtv1484) + +and _menhir_goto_expression : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_expression -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + match _menhir_s with + | MenhirState52 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1231 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | RBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1227 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1225 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _endpos_a_, _menhir_s, (a : 'tv_atom), _startpos_a_), _, _startpos__2_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _startpos = _startpos_a_ in + let _endpos = _endpos__4_ in + let _v : 'tv_atom = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos_a_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 499 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEindex (a, e)) ) +# 4997 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_atom _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1226)) : 'freshtv1228) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1229 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1230)) : 'freshtv1232) + | MenhirState54 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1237 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | BARBAR | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EQUAL | EVENT | EXTERNAL | GREATER | GREATEREQ | IN | LAYER | LESS | LESSEQ | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | UNEQUAL | WITH | XOR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1233 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 516 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPxor, e1, e2 )) ) +# 5054 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1234) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1235 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1236)) : 'freshtv1238) + | MenhirState56 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1241 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1239 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 522 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPtimes, e1, e2 )) ) +# 5079 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1240)) : 'freshtv1242) + | MenhirState58 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1245 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1243 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 523 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPdivide, e1, e2 )) ) +# 5097 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1244)) : 'freshtv1246) + | MenhirState60 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1251 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | BARBAR | BITAND | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EQUAL | EVENT | EXTERNAL | GREATER | GREATEREQ | IN | LAYER | LESS | LESSEQ | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SHL | SHR | SIGNATURE | THEN | TO | TRUSTED | TYPE | UNEQUAL | WITH | XOR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1247 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 519 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPshr, e1, e2)) ) +# 5129 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1248) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1249 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1250)) : 'freshtv1252) + | MenhirState62 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1257 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | BARBAR | BITAND | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EQUAL | EVENT | EXTERNAL | GREATER | GREATEREQ | IN | LAYER | LESS | LESSEQ | LET | LOGICAL | MINUS | OBJECT | PLUS | RBRACE | RBRACKET | RPAREN | SEMICOLON | SHL | SHR | SIGNATURE | THEN | TO | TRUSTED | TYPE | UNEQUAL | WITH | XOR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1253 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 520 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPplus, e1, e2 )) ) +# 5164 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1254) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1255 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1256)) : 'freshtv1258) + | MenhirState64 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1261 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1259 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 524 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPremainder, e1, e2 )) ) +# 5189 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1260)) : 'freshtv1262) + | MenhirState66 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1267 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | BARBAR | BITAND | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EQUAL | EVENT | EXTERNAL | GREATER | GREATEREQ | IN | LAYER | LESS | LESSEQ | LET | LOGICAL | MINUS | OBJECT | PLUS | RBRACE | RBRACKET | RPAREN | SEMICOLON | SHL | SHR | SIGNATURE | THEN | TO | TRUSTED | TYPE | UNEQUAL | WITH | XOR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1263 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _startpos__2_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 521 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPminus, e1, e2 )) ) +# 5217 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1264) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1265 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1266)) : 'freshtv1268) + | MenhirState68 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1273 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | BARBAR | BITAND | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EQUAL | EVENT | EXTERNAL | GREATER | GREATEREQ | IN | LAYER | LESS | LESSEQ | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SHL | SHR | SIGNATURE | THEN | TO | TRUSTED | TYPE | UNEQUAL | WITH | XOR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1269 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 518 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPshl, e1, e2 )) ) +# 5256 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1270) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1271 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1272)) : 'freshtv1274) + | MenhirState70 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1279 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | BARBAR | BITAND | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EQUAL | EVENT | EXTERNAL | GREATER | GREATEREQ | IN | LAYER | LESS | LESSEQ | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | UNEQUAL | WITH | XOR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1275 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 517 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPbitand, e1, e2 )) ) +# 5299 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1276) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1277 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1278)) : 'freshtv1280) + | MenhirState72 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1285 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EVENT | EXTERNAL | IN | LAYER | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1281 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 528 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPne, e1, e2 )) ) +# 5348 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1282) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1283 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1284)) : 'freshtv1286) + | MenhirState74 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1291 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | BARBAR | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EQUAL | EVENT | EXTERNAL | GREATER | GREATEREQ | IN | LAYER | LESS | LESSEQ | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | UNEQUAL | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1287 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 515 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPbitor, e1, e2 )) ) +# 5395 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1288) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1289 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1290)) : 'freshtv1292) + | MenhirState77 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1297 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EVENT | EXTERNAL | IN | LAYER | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1293 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 530 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPle, e1, e2 )) ) +# 5444 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1294) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1295 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1296)) : 'freshtv1298) + | MenhirState79 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1303 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EVENT | EXTERNAL | IN | LAYER | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1299 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 529 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPlt, e1, e2 )) ) +# 5493 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1300) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1301 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1302)) : 'freshtv1304) + | MenhirState81 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1309 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EVENT | EXTERNAL | IN | LAYER | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1305 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 532 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPge, e1, e2 )) ) +# 5542 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1306) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1307 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1308)) : 'freshtv1310) + | MenhirState83 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1315 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EVENT | EXTERNAL | IN | LAYER | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1311 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 531 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPgt, e1, e2 )) ) +# 5591 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1312) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1313 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1314)) : 'freshtv1316) + | MenhirState85 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1321 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EVENT | EXTERNAL | IN | LAYER | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1317 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 527 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPeq, e1, e2 )) ) +# 5640 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1318) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1319 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1320)) : 'freshtv1322) + | MenhirState87 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1327 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | COMMA | CONST | DO | ELSE | END | EOF | EVENT | EXTERNAL | IN | LAYER | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1323 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 526 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPor,e1, e2 )) ) +# 5705 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1324) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1325 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1326)) : 'freshtv1328) + | MenhirState89 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1333 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | ASSIGN | BAR | COMMA | CONST | DISJUNCTION | DO | ELSE | END | EOF | EVENT | EXTERNAL | IN | LAYER | LET | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | THEN | TO | TRUSTED | TYPE | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1329 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_expression = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 525 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEbin (OPand, e1, e2 )) ) +# 5768 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1330) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1331 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1332)) : 'freshtv1334) + | MenhirState47 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1337 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1335 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_) = _menhir_stack in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos_e_ in + let _v : 'tv_expression = let _endpos = _endpos_e_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 513 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEun (OPnot, e)) ) +# 5793 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1336)) : 'freshtv1338) + | MenhirState46 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1341 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1339 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_) = _menhir_stack in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos_e_ in + let _v : 'tv_expression = let _endpos = _endpos_e_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 514 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEun (OPbitnot, e)) ) +# 5811 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1340)) : 'freshtv1342) + | MenhirState37 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1357 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 5819 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | RBRACE | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1353 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 5865 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 5870 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_) = _menhir_stack in + let _2 = () in + let _v : 'tv_struct_field = let _endpos = _endpos_e_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 544 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _1, mkexp_ ~loc:_sloc e.p_expression_desc ) +# 5879 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1351) = _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_struct_field) = _v in + ((match _menhir_s with + | MenhirState96 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1345 * _menhir_state * 'tv_struct_fields) * _menhir_state) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_struct_field) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1343 * _menhir_state * 'tv_struct_fields) * _menhir_state) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_struct_field) : 'tv_struct_field) = _v in + ((let ((_menhir_stack, _menhir_s, (_1 : 'tv_struct_fields)), _) = _menhir_stack in + let _2 = () in + let _v : 'tv_struct_fields = +# 541 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 5900 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_struct_fields _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1344)) : 'freshtv1346) + | MenhirState35 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1349) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_struct_field) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1347) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_struct_field) : 'tv_struct_field) = _v in + ((let _v : 'tv_struct_fields = +# 540 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 5915 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_struct_fields _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1348)) : 'freshtv1350) + | _ -> + _menhir_fail ()) : 'freshtv1352)) : 'freshtv1354) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1355 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 5927 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1356)) : 'freshtv1358) + | MenhirState102 | MenhirState33 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1365 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | COMMA -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1359 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState102 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState102 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState102 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState102 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState102 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState102 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState102 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState102 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState102 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState102) : 'freshtv1360) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1361 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _endpos_e_, _menhir_s, (e : 'tv_expression), _startpos_e_) = _menhir_stack in + let _endpos = _endpos_e_ in + let _v : 'tv_comma_sep_expressions = let _endpos = _endpos_e_ in + let _symbolstartpos = _startpos_e_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 537 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc e.p_expression_desc ) +# 6012 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_comma_sep_expressions _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1362) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1363 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1364)) : 'freshtv1366) + | MenhirState32 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1369 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1367 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_) = _menhir_stack in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos_e_ in + let _v : 'tv_expression = let _endpos = _endpos_e_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 512 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEun (OPneg, e)) ) +# 6037 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1368)) : 'freshtv1370) + | MenhirState107 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1375 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1371 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BAR -> + _menhir_run110 _menhir_env (Obj.magic _menhir_stack) MenhirState109 + | IDENT _ | LBRACKET -> + _menhir_reduce208 _menhir_env (Obj.magic _menhir_stack) MenhirState109 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState109) : 'freshtv1372) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1373 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1374)) : 'freshtv1376) + | MenhirState118 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1381 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | THEN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1377 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState120 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState120 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState120 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState120 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState120) : 'freshtv1378) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1379 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1380)) : 'freshtv1382) + | MenhirState130 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1387 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6209 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | TO -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1383 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6251 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState132 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState132 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState132 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState132 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState132 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState132 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState132 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState132 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState132 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState132) : 'freshtv1384) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1385 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6289 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1386)) : 'freshtv1388) + | MenhirState132 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1393 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6298 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | DO -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1389 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6316 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState134 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState134 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState134 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState134 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState134) : 'freshtv1390) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1391 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6402 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1392)) : 'freshtv1394) + | MenhirState138 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1399 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6411 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | TO -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1395 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6453 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState140 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState140 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState140 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState140 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState140 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState140 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState140 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState140 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState140 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState140) : 'freshtv1396) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1397 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6491 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1398)) : 'freshtv1400) + | MenhirState140 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1409 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6500 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BAR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1405 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6510 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1401 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6520 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6526 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSIGN -> + _menhir_run129 _menhir_env (Obj.magic _menhir_stack) MenhirState143 + | EQUAL -> + _menhir_run128 _menhir_env (Obj.magic _menhir_stack) MenhirState143 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState143) : 'freshtv1402) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1403 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6548 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1404)) : 'freshtv1406) + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1407 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6595 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1408)) : 'freshtv1410) + | MenhirState144 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((('freshtv1415 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6604 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6608 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | DO -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((('freshtv1411 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6626 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6630 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState146 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState146 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState146 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState146 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState146) : 'freshtv1412) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((((('freshtv1413 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6716 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6720 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1414)) : 'freshtv1416) + | MenhirState150 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1421 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6729 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | TO -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1417 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6771 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState152 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState152 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState152 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState152 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState152 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState152 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState152 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState152 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState152 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState152) : 'freshtv1418) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1419 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6809 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1420)) : 'freshtv1422) + | MenhirState152 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1427 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6818 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | DO -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1423 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6836 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState154 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState154 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState154 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState154 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState154) : 'freshtv1424) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv1425 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 6922 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1426)) : 'freshtv1428) + | MenhirState156 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1433 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | BAR | COMMA | ELSE | END | IN | LET | RBRACE | RBRACKET | RPAREN | SEMICOLON | THEN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1429 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_) = _menhir_stack in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos_e_ in + let _v : 'tv_command_core = let _endpos = _endpos_e_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 559 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCemit (mkexp_ ~loc:_sloc e.p_expression_desc) ) +# 6981 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1430) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1431 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1432)) : 'freshtv1434) + | MenhirState501 | MenhirState491 | MenhirState30 | MenhirState201 | MenhirState114 | MenhirState190 | MenhirState117 | MenhirState187 | MenhirState120 | MenhirState121 | MenhirState183 | MenhirState124 | MenhirState134 | MenhirState146 | MenhirState178 | MenhirState176 | MenhirState174 | MenhirState154 | MenhirState158 | MenhirState170 | MenhirState159 | MenhirState160 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1441 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSIGN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1435 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState162 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState162 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState162 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState162 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState162 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState162 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState162 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState162 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState162 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState162) : 'freshtv1436) + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | BAR | COMMA | ELSE | END | IN | LET | RBRACE | RBRACKET | RPAREN | SEMICOLON | THEN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1437 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _endpos_e_, _menhir_s, (e : 'tv_expression), _startpos_e_) = _menhir_stack in + let _startpos = _startpos_e_ in + let _endpos = _endpos_e_ in + let _v : 'tv_command_core = +# 548 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCyield e ) +# 7070 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1438) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1439 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1440)) : 'freshtv1442) + | MenhirState162 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1447 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | BAR | COMMA | ELSE | END | IN | LET | RBRACE | RBRACKET | RPAREN | SEMICOLON | THEN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1443 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_e1_, _menhir_s, (e1 : 'tv_expression), _startpos_e1_), _endpos_e2_, _, (e2 : 'tv_expression), _startpos_e2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_e1_ in + let _endpos = _endpos_e2_ in + let _v : 'tv_command_core = let _endpos = _endpos_e2_ in + let _symbolstartpos = _startpos_e1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 553 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( (PCstore (mkexp_ ~loc:_sloc e1.p_expression_desc, mkexp_ ~loc:_sloc e2.p_expression_desc)) ) +# 7135 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1444) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1445 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1446)) : 'freshtv1448) + | MenhirState228 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1453 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | RBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1449 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_stack = (_menhir_stack, _endpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState230 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState230 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState230 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState230 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState230 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState230 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState230) : 'freshtv1450) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1451 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1452)) : 'freshtv1454) + | MenhirState445 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1465) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7223 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1461) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7269 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__2_, (_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7274 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__2_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_) = _menhir_stack in + let _3 = () in + let _1 = () in + let _endpos = _endpos_e_ in + let _v : 'tv_const_declaration = +# 232 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( Hashtbl.add constant_table _2 e.p_expression_desc ) +# 7282 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1459) = _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_const_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1457 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_const_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1455 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let ((_2 : 'tv_const_declaration) : 'tv_const_declaration) = _v in + ((let (_menhir_stack, _endpos__1_, (_1 : 'tv_declarations)) = _menhir_stack in + let _endpos = _endpos__2_ in + let _v : 'tv_declarations = +# 220 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _1 ) +# 7301 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_declarations _menhir_env _menhir_stack _endpos _v) : 'freshtv1456)) : 'freshtv1458)) : 'freshtv1460)) : 'freshtv1462) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1463) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7311 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1464)) : 'freshtv1466) + | MenhirState497 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1471 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7320 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BARBAR -> + _menhir_run74 _menhir_env (Obj.magic _menhir_stack) + | BITAND -> + _menhir_run70 _menhir_env (Obj.magic _menhir_stack) + | CONJUNCTION -> + _menhir_run89 _menhir_env (Obj.magic _menhir_stack) + | DISJUNCTION -> + _menhir_run87 _menhir_env (Obj.magic _menhir_stack) + | EQUAL -> + _menhir_run85 _menhir_env (Obj.magic _menhir_stack) + | GREATER -> + _menhir_run83 _menhir_env (Obj.magic _menhir_stack) + | GREATEREQ -> + _menhir_run81 _menhir_env (Obj.magic _menhir_stack) + | LESS -> + _menhir_run79 _menhir_env (Obj.magic _menhir_stack) + | LESSEQ -> + _menhir_run77 _menhir_env (Obj.magic _menhir_stack) + | MINUS -> + _menhir_run66 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MOD -> + _menhir_run64 _menhir_env (Obj.magic _menhir_stack) + | PLUS -> + _menhir_run62 _menhir_env (Obj.magic _menhir_stack) + | SHL -> + _menhir_run68 _menhir_env (Obj.magic _menhir_stack) + | SHR -> + _menhir_run60 _menhir_env (Obj.magic _menhir_stack) + | SLASH -> + _menhir_run58 _menhir_env (Obj.magic _menhir_stack) + | STAR -> + _menhir_run56 _menhir_env (Obj.magic _menhir_stack) + | UNEQUAL -> + _menhir_run72 _menhir_env (Obj.magic _menhir_stack) + | XOR -> + _menhir_run54 _menhir_env (Obj.magic _menhir_stack) + | LET | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1467 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7366 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_annotations)), _, (_3 : 'tv_method_kind)), _endpos__4_, (_4 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7371 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__4_), _), _endpos__6_, _, (_6 : 'tv_type_expression), _startpos__6_), _), _endpos_e_, _, (e : 'tv_expression), _startpos_e_) = _menhir_stack in + let _7 = () in + let _5 = () in + let _1 = () in + let _v : 'tv_object_field_or_method = let _endpos = _endpos_e_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 682 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( (* annotations are not used and only one (GHOST) method_kind is + acceptable, just there to allow parser to factor with + the next rule *) + let is_ghost = match _3 with + | MKnormal -> false + | MKghost -> true + | MKconstghost -> + print_endline ("Warning: object fields cannot be declared `const ghost': " ^ + _4 ^ " marked as `ghost'."); + true + | k -> + print_endline ("Warning: object fields cannot be declared `" ^ + string_of_method_kind k ^ "': " ^ _4 ^ " marked as normal."); + false + in `Left (_4, _6, (mkexp_ ~loc:_sloc e.p_expression_desc) , is_ghost) ) +# 7396 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_field_or_method _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1468) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1469 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7406 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1470)) : 'freshtv1472) + | _ -> + _menhir_fail () + +and _menhir_goto_method_kind : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_method_kind -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + match _menhir_s with + | MenhirState273 | MenhirState287 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1215 * _menhir_state * 'tv_method_kind) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1211 * _menhir_state * 'tv_method_kind) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7430 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1207 * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7442 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState291 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState291 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState291 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState291 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState291 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState291 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState291) : 'freshtv1208) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1209 * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7470 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1210)) : 'freshtv1212) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1213 * _menhir_state * 'tv_method_kind) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1214)) : 'freshtv1216) + | MenhirState475 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1223 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1219 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7494 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1217 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7506 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState494 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState495 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState495 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState495 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState495 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState495 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState495 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState495) : 'freshtv1218) + | IDENT _v -> + _menhir_run488 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState494 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run477 _menhir_env (Obj.magic _menhir_stack) MenhirState494 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState494) : 'freshtv1220) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1221 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1222)) : 'freshtv1224) + | _ -> + _menhir_fail () + +and _menhir_run474 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + _menhir_run5 _menhir_env (Obj.magic _menhir_stack) MenhirState474 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | CONST | CONSTRUCTOR | GHOST | IDENT _ | LOGICAL | REFINED -> + _menhir_reduce19 _menhir_env (Obj.magic _menhir_stack) MenhirState474 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState474 + +and _menhir_goto_base_slots : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_base_slots -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1205 * Lexing.position) * _menhir_state * 'tv_base_slots) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COMMA -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1197 * _menhir_state * 'tv_base_slots) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1195) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState464 in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7584 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + _menhir_run458 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1193 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7600 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1194)) : 'freshtv1196) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState464) : 'freshtv1198) + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1201 * Lexing.position) * _menhir_state * 'tv_base_slots) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1199 * Lexing.position) * _menhir_state * 'tv_base_slots) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + ((let ((_menhir_stack, _startpos__1_), _, (_2 : 'tv_base_slots)) = _menhir_stack in + let _3 = () in + let _1 = () in + let _v : 'tv_base_layer_signature = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 664 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayersign ~loc:_sloc (PLSconstr (List.rev _2)) ) +# 7625 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_base_layer_signature _menhir_env _menhir_stack _v) : 'freshtv1200)) : 'freshtv1202) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1203 * Lexing.position) * _menhir_state * 'tv_base_slots) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1204)) : 'freshtv1206) + +and _menhir_goto_layer_signature_fields : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_layer_signature_fields -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1191 * _menhir_state * Lexing.position) * _menhir_state * 'tv_layer_signature_fields) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1189 * _menhir_state * 'tv_layer_signature_fields) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState331 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run328 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState332 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RBRACE -> + _menhir_reduce214 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState332) : 'freshtv1190) + | RBRACE -> + _menhir_reduce213 _menhir_env (Obj.magic _menhir_stack) MenhirState331 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState331) : 'freshtv1192) + +and _menhir_goto_object_signature_declaration : _menhir_env -> 'ttv_tail -> Lexing.position -> 'tv_object_signature_declaration -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v _startpos -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1187) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_object_signature_declaration) = _v in + let (_startpos : Lexing.position) = _startpos in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1185) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let ((_1 : 'tv_object_signature_declaration) : 'tv_object_signature_declaration) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _endpos = _endpos__1_ in + let _v : 'tv_declaration = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 225 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( fst _1, mkdecl ~loc:_sloc (PDsignature (snd _1)) ) +# 7686 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv1186)) : 'freshtv1188) + +and _menhir_run300 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | GHOST -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1171 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1165 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7710 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1163 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + let ((_4 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7720 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7724 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__4_ : Lexing.position) = _startpos in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_object_signature_expr), _startpos__1_), _startpos__3_) = _menhir_stack in + let _3 = () in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__4_ in + let _v : 'tv_object_signature_expr = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 361 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mksig ~loc:_sloc (PSghostize (_1, [_4])) ) +# 7738 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_expr _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1164)) : 'freshtv1166) + | LBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1167 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run303 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState312 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState312) : 'freshtv1168) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1169 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1170)) : 'freshtv1172) + | LOGICAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1181 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1175 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7778 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1173 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + let ((_4 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7788 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7792 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__4_ : Lexing.position) = _startpos in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_object_signature_expr), _startpos__1_), _endpos__3_, _startpos__3_) = _menhir_stack in + let _3 = () in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__4_ in + let _v : 'tv_object_signature_expr = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 363 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mksig ~loc:_sloc (PSlogicize (_1, [_4])) ) +# 7806 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_expr _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1174)) : 'freshtv1176) + | LBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1177 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run303 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState302 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState302) : 'freshtv1178) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1179 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _, _), _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1180)) : 'freshtv1182) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1183 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1184) + +and _menhir_run316 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _startpos -> + let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1157 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7851 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1155 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + let ((_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7861 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 7865 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__3_ : Lexing.position) = _startpos in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_object_signature_expr), _startpos__1_), _startpos__2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_object_signature_expr = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 365 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mksig ~loc:_sloc (PSminus (_1, [_3])) ) +# 7878 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_expr _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1156)) : 'freshtv1158) + | LBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1159 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run303 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState317 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState317) : 'freshtv1160) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1161 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1162) + +and _menhir_goto_object_expression : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_object_expression -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + match _menhir_s with + | MenhirState355 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1123 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLONGREATER -> + _menhir_run361 _menhir_env (Obj.magic _menhir_stack) + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1119 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1117 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_object_expression), _startpos__2_) = _menhir_stack in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_object_expression = +# 735 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 ) +# 7931 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1118)) : 'freshtv1120) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1121 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1122)) : 'freshtv1124) + | MenhirState367 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1129 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 7946 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLONGREATER -> + _menhir_run361 _menhir_env (Obj.magic _menhir_stack) + | RBRACE | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1125 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 7958 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 7963 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _endpos__4_), _endpos__6_, _, (_6 : 'tv_object_expression), _startpos__6_) = _menhir_stack in + let _5 = () in + let _4 = () in + let _2 = () in + let _1 = () in + let _v : 'tv_layer_obj_inst = let _endpos = _endpos__6_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 784 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkobjinst ~loc:_sloc (POexternal ((CONaddress _3), _6)) ) +# 7975 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_obj_inst _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1126) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1127 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 7985 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1128)) : 'freshtv1130) + | MenhirState371 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1135 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 7994 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLONGREATER -> + _menhir_run361 _menhir_env (Obj.magic _menhir_stack) + | RBRACE | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1131 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8006 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (((((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8011 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _endpos__4_), _endpos__6_, _, (_6 : 'tv_object_expression), _startpos__6_) = _menhir_stack in + let _5 = () in + let _4 = () in + let _2 = () in + let _1 = () in + let _v : 'tv_layer_obj_inst = let _endpos = _endpos__6_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 783 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkobjinst ~loc:_sloc (POexternal ((CONaddress _3), _6)) ) +# 8023 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_obj_inst _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1132) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1133 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8033 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1134)) : 'freshtv1136) + | MenhirState354 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1141 * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLONGREATER -> + _menhir_run361 _menhir_env (Obj.magic _menhir_stack) + | RBRACE | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1137 * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_object_expression), _startpos__1_) = _menhir_stack in + let _v : 'tv_layer_obj_inst = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 782 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkobjinst ~loc:_sloc (POinternal _1) ) +# 8055 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_obj_inst _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1138) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1139 * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1140)) : 'freshtv1142) + | MenhirState468 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1147 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8070 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLONGREATER -> + _menhir_run361 _menhir_env (Obj.magic _menhir_stack) + | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1143 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8082 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _endpos__1_, (_1 : 'tv_opt_logical_or_trusted), _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8087 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _endpos__5_, _, (_5 : 'tv_object_expression), _startpos__5_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__5_ in + let _v : 'tv_object_declaration = let _endpos = _endpos__5_ in + let _symbolstartpos = if _startpos__1_ != _endpos__1_ then + _startpos__1_ + else + _startpos__2_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 657 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3, { pObjectType = None; pObjectDesc = _5; pObjectLoc = (make_loc _sloc) } ) +# 8102 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_declaration _menhir_env _menhir_stack _endpos _v _startpos) : 'freshtv1144) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1145 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8112 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1146)) : 'freshtv1148) + | MenhirState507 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1153 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8121 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLONGREATER -> + _menhir_run361 _menhir_env (Obj.magic _menhir_stack) + | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1149 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8133 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((((_menhir_stack, _endpos__1_, (_1 : 'tv_opt_logical_or_trusted), _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8138 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), (_4 : 'tv_base_layer_signature)), _endpos__6_, _, (_6 : 'tv_object_signature_expr), _startpos__6_), _endpos__8_, _, (_8 : 'tv_object_expression), _startpos__8_) = _menhir_stack in + let _7 = () in + let _5 = () in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__8_ in + let _v : 'tv_object_declaration = let _endpos = _endpos__8_ in + let _symbolstartpos = if _startpos__1_ != _endpos__1_ then + _startpos__1_ + else + _startpos__2_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 654 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3, { pObjectType = Some { pObjectBase = _4; pObjectSignature = _6; pObjectTypLoc = (make_loc _sloc) }; + pObjectDesc = _8; pObjectLoc = (make_loc _sloc) } ) +# 8155 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_declaration _menhir_env _menhir_stack _endpos _v _startpos) : 'freshtv1150) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv1151 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8165 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1152)) : 'freshtv1154) + | _ -> + _menhir_fail () + +and _menhir_goto_layer_type : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_layer_type -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v -> + match _menhir_s with + | MenhirState341 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1111) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_layer_type) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1109) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_2 : 'tv_layer_type) : 'tv_layer_type) = _v in + ((let _1 = () in + let _v : 'tv_layer_signature_annotation = +# 753 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( Some _2 ) +# 8190 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_signature_annotation _menhir_env _menhir_stack _v) : 'freshtv1110)) : 'freshtv1112) + | MenhirState391 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1115 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_layer_type) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1113 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_layer_type) : 'tv_layer_type) = _v in + ((let (_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_layer_expression), _startpos__1_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_layer_expression = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 764 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayer ~loc:_sloc (PLrelax (_1, _3)) ) +# 8214 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1114)) : 'freshtv1116) + | _ -> + _menhir_fail () + +and _menhir_goto_declaration : _menhir_env -> 'ttv_tail -> Lexing.position -> 'tv_declaration -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1107 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1105 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let ((_2 : 'tv_declaration) : 'tv_declaration) = _v in + ((let (_menhir_stack, _endpos__1_, (_1 : 'tv_declarations)) = _menhir_stack in + let _endpos = _endpos__2_ in + let _v : 'tv_declarations = +# 218 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 :: _1 ) +# 8235 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_declarations _menhir_env _menhir_stack _endpos _v) : 'freshtv1106)) : 'freshtv1108) + +and _menhir_run477 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run479 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState477 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1103 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState477 in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1101 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _2 = () in + let _1 = () in + let _v : 'tv_method_param = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 717 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( ["()", Some (mkfotyp ~loc:_sloc (PTbuiltin Tunit))] ) +# 8266 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_param _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1102)) : 'freshtv1104) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState477 + +and _menhir_run488 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8277 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1099) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8288 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8292 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _v : 'tv_method_param = +# 716 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1, None] ) +# 8298 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_param _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1100) + +and _menhir_goto_external_declaration : _menhir_env -> 'ttv_tail -> Lexing.position -> 'tv_external_declaration -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1097 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_external_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1095 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let ((_2 : 'tv_external_declaration) : 'tv_external_declaration) = _v in + ((let (_menhir_stack, _endpos__1_, (_1 : 'tv_declarations)) = _menhir_stack in + let _endpos = _endpos__2_ in + let _v : 'tv_declarations = +# 221 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 :: _1 ) +# 8317 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_declarations _menhir_env _menhir_stack _endpos _v) : 'freshtv1096)) : 'freshtv1098) + +and _menhir_goto_layer_signature_annotation : _menhir_env -> 'ttv_tail -> 'tv_layer_signature_annotation -> 'ttv_return = + fun _menhir_env _menhir_stack _v -> + let _menhir_stack = (_menhir_stack, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1093 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8328 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * 'tv_layer_signature_annotation) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1089 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8338 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * 'tv_layer_signature_annotation) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run382 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState350 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run352 _menhir_env (Obj.magic _menhir_stack) MenhirState350 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run351 _menhir_env (Obj.magic _menhir_stack) MenhirState350 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState350) : 'freshtv1090) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv1091 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8360 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * 'tv_layer_signature_annotation) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1092)) : 'freshtv1094) + +and _menhir_run342 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run337 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState342 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run326 _menhir_env (Obj.magic _menhir_stack) MenhirState342 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1087 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState342 in + ((let _menhir_stack = (_menhir_stack, _endpos, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run337 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState343 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run326 _menhir_env (Obj.magic _menhir_stack) MenhirState343 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState343) : 'freshtv1088) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState342 + +and _menhir_goto_constructor_declarations : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_constructor_declarations -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1085 * _menhir_state * 'tv_opt_bar) * Lexing.position * _menhir_state * 'tv_constructor_declarations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | BAR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1079 * Lexing.position * _menhir_state * 'tv_constructor_declarations) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run255 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState266 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState266) : 'freshtv1080) + | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1081 * _menhir_state * 'tv_opt_bar) * Lexing.position * _menhir_state * 'tv_constructor_declarations) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, (_1 : 'tv_opt_bar)), _endpos__2_, _, (_2 : 'tv_constructor_declarations)) = _menhir_stack in + let _endpos = _endpos__2_ in + let _v : 'tv_type_definition = +# 280 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PTbranches (List.rev _2) ) +# 8425 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_definition _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv1082) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1083 * _menhir_state * 'tv_opt_bar) * Lexing.position * _menhir_state * 'tv_constructor_declarations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1084)) : 'freshtv1086) + +and _menhir_run244 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8439 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1075 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8451 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState245 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState245 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState245 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState245 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState245 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState245 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState245) : 'freshtv1076) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1077 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8479 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1078) + +and _menhir_goto_annotation : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_annotation -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + match _menhir_s with + | MenhirState211 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1069 * _menhir_state * 'tv_annotations_plus)) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_annotation) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1067 * _menhir_state * 'tv_annotations_plus)) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_annotation) : 'tv_annotation) = _v in + ((let (_menhir_stack, _menhir_s, (_1 : 'tv_annotations_plus)) = _menhir_stack in + let _2 = () in + let _v : 'tv_annotations_plus = +# 794 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 8501 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotations_plus _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1068)) : 'freshtv1070) + | MenhirState6 | MenhirState11 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1073) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_annotation) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1071) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_annotation) : 'tv_annotation) = _v in + ((let _v : 'tv_annotations_plus = +# 793 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 8516 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotations_plus _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1072)) : 'freshtv1074) + | _ -> + _menhir_fail () + +and _menhir_reduce10 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _v : 'tv_annotation_arguments = +# 807 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [] ) +# 8527 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation_arguments _menhir_env _menhir_stack _menhir_s _v + +and _menhir_run9 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8534 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1065) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8545 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8549 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _v : 'tv_annotation_arguments = let _endpos = _endpos__1_ in + let _startpos = _startpos__1_ in + let _loc = (_startpos, _endpos) in + +# 809 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [PAexpr (mkcmd ~loc:_loc (PCyield (mkexp_ ~loc:_loc (PEconstant (CONuint (int_of_string(_1)))))))] ) +# 8558 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation_arguments _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1066) + +and _menhir_run10 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8565 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1063) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8576 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8580 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _v : 'tv_annotation_arguments = +# 813 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [PAclause (_1, [])] ) +# 8586 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation_arguments _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1064) + +and _menhir_run11 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + _menhir_run30 _menhir_env (Obj.magic _menhir_stack) MenhirState11 + | IDENT _v -> + _menhir_run20 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState11 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run14 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState11 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run13 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState11 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1061 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState11 in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1059 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _2 = () in + let _1 = () in + let _v : 'tv_annotation_arguments = +# 814 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [] ) +# 8620 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation_arguments _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1060)) : 'freshtv1062) + | STRING _v -> + _menhir_run8 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState11 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run7 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState11 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState11 + +and _menhir_run21 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8635 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1057) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8646 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8650 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _v : 'tv_annotation_arguments = let _endpos = _endpos__1_ in + let _startpos = _startpos__1_ in + let _loc = (_startpos, _endpos) in + +# 808 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [PAexpr (mkcmd ~loc:_loc (PCyield (mkexp_ ~loc:_loc (PEconstant (CONint (int_of_string(_1)))))))] ) +# 8659 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation_arguments _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1058) + +and _menhir_run22 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1053 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | INT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1041 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8684 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1037 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8696 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1035 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8704 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8710 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _v : 'tv_annotation_arguments = let _endpos = _endpos__4_ in + let _startpos = _startpos__1_ in + let _loc = (_startpos, _endpos) in + +# 810 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [PAexpr (mkcmd ~loc:_loc (PCyield (mkexp_ ~loc:_loc (PEconstant (CONaddress _3)))))] ) +# 8721 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation_arguments _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1036)) : 'freshtv1038) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1039 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8731 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _), _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1040)) : 'freshtv1042) + | UINT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1049 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8742 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1045 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8754 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1043 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8762 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8768 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _v : 'tv_annotation_arguments = let _endpos = _endpos__4_ in + let _startpos = _startpos__1_ in + let _loc = (_startpos, _endpos) in + +# 811 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [PAexpr (mkcmd ~loc:_loc (PCyield (mkexp_ ~loc:_loc (PEconstant (CONaddress _3)))))] ) +# 8779 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation_arguments _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1044)) : 'freshtv1046) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv1047 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 8789 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _), _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1048)) : 'freshtv1050) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1051 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1052)) : 'freshtv1054) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1055 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1056) + +and _menhir_run28 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8811 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1033) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8822 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8826 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _v : 'tv_annotation_arguments = +# 812 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [PAclause (_1, [])] ) +# 8832 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation_arguments _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1034) + +and _menhir_run107 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState107 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState107 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState107 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState107 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState107 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState107 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState107 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState107 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState107 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState107 + +and _menhir_run115 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1029 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8878 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1025 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8890 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState117 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState117 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState117 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState117 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState117) : 'freshtv1026) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1027 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 8948 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1028)) : 'freshtv1030) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1031 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1032) + +and _menhir_run118 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState118 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState118 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState118 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState118 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState118 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState118 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState118 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState118 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState118 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState118 + +and _menhir_run121 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState121 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState121 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run122 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState121 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState121 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState121 + +and _menhir_run125 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + _menhir_run5 _menhir_env (Obj.magic _menhir_stack) MenhirState125 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _ -> + _menhir_reduce19 _menhir_env (Obj.magic _menhir_stack) MenhirState125 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState125 + +and _menhir_run135 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + _menhir_run5 _menhir_env (Obj.magic _menhir_stack) MenhirState135 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _ -> + _menhir_reduce19 _menhir_env (Obj.magic _menhir_stack) MenhirState135 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState135 + +and _menhir_run147 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + _menhir_run5 _menhir_env (Obj.magic _menhir_stack) MenhirState147 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _ -> + _menhir_reduce19 _menhir_env (Obj.magic _menhir_stack) MenhirState147 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState147 + +and _menhir_run155 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1023) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_command_core = +# 554 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PCfail ) +# 9095 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_command_core _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv1024) + +and _menhir_run156 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState156 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState156 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState156 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState156 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState156 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState156 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState156 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState156 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState156 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState156 + +and _menhir_run158 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState158 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState158 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run122 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState158 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState158 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState158 + +and _menhir_run159 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState159 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState159 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState159 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState159 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState159 + +and _menhir_run160 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState160 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState160 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run122 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState160 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState160 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState160 + +and _menhir_goto_opt_type_annotation : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_opt_type_annotation -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + match _menhir_s with + | MenhirState479 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1009 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9284 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_opt_type_annotation) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv1007 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9290 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_opt_type_annotation) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9295 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_), _, (_2 : 'tv_opt_type_annotation)) = _menhir_stack in + let _v : 'tv_method_parameter = +# 725 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _1, _2 ) +# 9300 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1005) = _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_method_parameter) = _v in + ((match _menhir_s with + | MenhirState485 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv999 * _menhir_state * 'tv_method_parameters)) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_method_parameter) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv997 * _menhir_state * 'tv_method_parameters)) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_method_parameter) : 'tv_method_parameter) = _v in + ((let (_menhir_stack, _menhir_s, (_1 : 'tv_method_parameters)) = _menhir_stack in + let _2 = () in + let _v : 'tv_method_parameters = +# 722 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 9321 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_parameters _menhir_env _menhir_stack _menhir_s _v) : 'freshtv998)) : 'freshtv1000) + | MenhirState477 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1003) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_method_parameter) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1001) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_method_parameter) : 'tv_method_parameter) = _v in + ((let _v : 'tv_method_parameters = +# 721 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 9336 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_parameters _menhir_env _menhir_stack _menhir_s _v) : 'freshtv1002)) : 'freshtv1004) + | _ -> + _menhir_fail ()) : 'freshtv1006)) : 'freshtv1008)) : 'freshtv1010) + | MenhirState489 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1015 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1011 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState491 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState491 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState491 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState491 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState491) : 'freshtv1012) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv1013 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1014)) : 'freshtv1016) + | MenhirState499 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1021 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9411 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1017 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9421 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState501 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState501 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState501 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState501 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState501) : 'freshtv1018) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv1019 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9479 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv1020)) : 'freshtv1022) + | _ -> + _menhir_fail () + +and _menhir_goto_indexed_opt : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_indexed_opt -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv995 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9493 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_indexed_opt) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv991 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9503 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_indexed_opt) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv989 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9511 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_indexed_opt) = Obj.magic _menhir_stack in + let (_endpos__6_ : Lexing.position) = _endpos in + ((let ((((_menhir_stack, _startpos__1_), _endpos__2_, (_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9517 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__2_), _endpos__4_, _, (_4 : 'tv_type_expression), _startpos__4_), _, (_5 : 'tv_indexed_opt)) = _menhir_stack in + let _6 = () in + let _3 = () in + let _1 = () in + let _endpos = _endpos__6_ in + let _v : 'tv_event_parameter = +# 344 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( (_2, _4, _5) ) +# 9526 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv987) = _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_event_parameter) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv985 * Lexing.position * 'tv_event_parameters) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_event_parameter) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv983 * Lexing.position * 'tv_event_parameters) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let ((_2 : 'tv_event_parameter) : 'tv_event_parameter) = _v in + ((let (_menhir_stack, _endpos__1_, (_1 : 'tv_event_parameters)) = _menhir_stack in + let _endpos = _endpos__2_ in + let _v : 'tv_event_parameters = +# 334 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 :: _1 ) +# 9545 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_event_parameters _menhir_env _menhir_stack _endpos _v) : 'freshtv984)) : 'freshtv986)) : 'freshtv988)) : 'freshtv990)) : 'freshtv992) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv993 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9555 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state * 'tv_indexed_opt) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv994)) : 'freshtv996) + +and _menhir_run128 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv981) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + ((let _1 = () in + let _v : 'tv_eq_or_assign = +# 593 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( () ) +# 9570 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_eq_or_assign _menhir_env _menhir_stack _menhir_s _v) : 'freshtv982) + +and _menhir_run129 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv979) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + ((let _1 = () in + let _v : 'tv_eq_or_assign = +# 594 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( () ) +# 9584 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_eq_or_assign _menhir_env _menhir_stack _menhir_s _v) : 'freshtv980) + +and _menhir_goto_object_signature_field : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_object_signature_field -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + match _menhir_s with + | MenhirState287 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv973 * _menhir_state * 'tv_object_signature_fields) * _menhir_state) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_object_signature_field) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv971 * _menhir_state * 'tv_object_signature_fields) * _menhir_state) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_object_signature_field) : 'tv_object_signature_field) = _v in + ((let ((_menhir_stack, _menhir_s, (_1 : 'tv_object_signature_fields)), _) = _menhir_stack in + let _2 = () in + let _v : 'tv_object_signature_fields = +# 369 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 9605 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_fields _menhir_env _menhir_stack _menhir_s _v) : 'freshtv972)) : 'freshtv974) + | MenhirState273 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv977) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_object_signature_field) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv975) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_object_signature_field) : 'tv_object_signature_field) = _v in + ((let _v : 'tv_object_signature_fields = +# 368 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 9620 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_fields _menhir_env _menhir_stack _menhir_s _v) : 'freshtv976)) : 'freshtv978) + | _ -> + _menhir_fail () + +and _menhir_goto_constructor_parameters : _menhir_env -> 'ttv_tail -> Lexing.position -> 'tv_constructor_parameters -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v -> + let _menhir_stack = (_menhir_stack, _endpos, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv969 * Lexing.position * 'tv_constructor_parameters) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | LPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv959) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run194 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState259 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState259) : 'freshtv960) + | BAR | CONST | EOF | EVENT | EXTERNAL | LAYER | LBRACKET | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv965 * Lexing.position * 'tv_constructor_parameters) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _endpos__1_, (_1 : 'tv_constructor_parameters)) = _menhir_stack in + let _endpos = _endpos__1_ in + let _v : 'tv_constructor_params = +# 304 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( List.fold_left (fun a b -> List.rev_append b a) [] _1 ) +# 9656 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv963) = _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_constructor_params) = _v in + ((let _menhir_stack = (_menhir_stack, _endpos, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv961 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9667 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * 'tv_constructor_params) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + _menhir_run5 _menhir_env (Obj.magic _menhir_stack) MenhirState256 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BAR | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + _menhir_reduce19 _menhir_env (Obj.magic _menhir_stack) MenhirState256 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState256) : 'freshtv962)) : 'freshtv964)) : 'freshtv966) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv967 * Lexing.position * 'tv_constructor_parameters) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv968)) : 'freshtv970) + +and _menhir_goto_field_declarations : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_field_declarations -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv957 * _menhir_state * Lexing.position) * _menhir_state * 'tv_field_declarations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv955 * _menhir_state * 'tv_field_declarations) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState247 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run244 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState248 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RBRACE -> + _menhir_reduce214 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState248) : 'freshtv956) + | RBRACE -> + _menhir_reduce213 _menhir_env (Obj.magic _menhir_stack) MenhirState247 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState247) : 'freshtv958) + +and _menhir_goto_type_declaration : _menhir_env -> 'ttv_tail -> Lexing.position -> 'tv_type_declaration -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v _startpos -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv953) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_type_declaration) = _v in + let (_startpos : Lexing.position) = _startpos in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv951) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let ((_1 : 'tv_type_declaration) : 'tv_type_declaration) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _endpos = _endpos__1_ in + let _v : 'tv_declaration = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 224 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( fst _1, mkdecl ~loc:_sloc (PDtype (snd _1)) ) +# 9737 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv952)) : 'freshtv954) + +and _menhir_run233 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState233 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState233 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState233 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState233 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState233 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState233 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState233 + +and _menhir_fail : unit -> 'a = + fun () -> + Printf.fprintf stderr "Internal failure -- please contact the parser generator's developers.\n%!"; + assert false + +and _menhir_run428 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9772 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv949) = Obj.magic _menhir_stack in + ((let (_, _endpos) = Obj.magic _menhir_stack in + let _v : 'tv_event_parameters = +# 333 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [] ) +# 9783 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_event_parameters _menhir_env _menhir_stack _endpos _v) : 'freshtv950) + +and _menhir_run255 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9790 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv947) = Obj.magic _menhir_stack in + ((let (_, _endpos) = Obj.magic _menhir_stack in + let _v : 'tv_constructor_parameters = +# 307 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [] ) +# 9801 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_constructor_parameters _menhir_env _menhir_stack _endpos _v) : 'freshtv948) + +and _menhir_run112 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv943 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_stack = (_menhir_stack, _endpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | DOUBLEARROW -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv939 * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState114 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState114 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState114 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState114 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState114) : 'freshtv940) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv941 * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv942)) : 'freshtv944) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv945 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv946) + +and _menhir_run193 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9889 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLONCOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv935) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState193 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv931 * _menhir_state) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9911 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv929 * _menhir_state) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let ((_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9921 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 9925 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__2_ : Lexing.position) = _startpos in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + let _1 = () in + let _v : 'tv_match_pattern_tail = +# 628 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PPTcons _2 ) +# 9933 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_match_pattern_tail _menhir_env _menhir_stack _menhir_s _v) : 'freshtv930)) : 'freshtv932) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv933 * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv934)) : 'freshtv936) + | IDENT _v -> + _menhir_run194 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState193 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DOUBLEARROW -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv937) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState193 in + ((let _v : 'tv_match_pattern_tail = +# 627 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PPTother [] ) +# 9952 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_match_pattern_tail _menhir_env _menhir_stack _menhir_s _v) : 'freshtv938) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState193 + +and _menhir_goto_nonempty_list_atom_ : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_nonempty_list_atom_ -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + match _menhir_s with + | MenhirState501 | MenhirState497 | MenhirState491 | MenhirState445 | MenhirState228 | MenhirState30 | MenhirState201 | MenhirState114 | MenhirState190 | MenhirState117 | MenhirState187 | MenhirState120 | MenhirState121 | MenhirState183 | MenhirState124 | MenhirState134 | MenhirState146 | MenhirState178 | MenhirState176 | MenhirState174 | MenhirState154 | MenhirState158 | MenhirState170 | MenhirState159 | MenhirState162 | MenhirState160 | MenhirState156 | MenhirState152 | MenhirState150 | MenhirState144 | MenhirState140 | MenhirState138 | MenhirState132 | MenhirState130 | MenhirState118 | MenhirState107 | MenhirState32 | MenhirState102 | MenhirState33 | MenhirState37 | MenhirState46 | MenhirState89 | MenhirState87 | MenhirState85 | MenhirState83 | MenhirState81 | MenhirState79 | MenhirState77 | MenhirState74 | MenhirState72 | MenhirState70 | MenhirState68 | MenhirState66 | MenhirState64 | MenhirState62 | MenhirState60 | MenhirState58 | MenhirState56 | MenhirState54 | MenhirState52 | MenhirState47 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv919 * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | DOT -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv913 * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState49 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState49 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState49 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState49 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState49 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState49) : 'freshtv914) + | ASSIGN | BAR | BARBAR | BITAND | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | ELSE | END | EOF | EQUAL | EVENT | EXTERNAL | GREATER | GREATEREQ | IN | LAYER | LESS | LESSEQ | LET | LOGICAL | MINUS | MOD | OBJECT | PLUS | RBRACE | RBRACKET | RPAREN | SEMICOLON | SHL | SHR | SIGNATURE | SLASH | STAR | THEN | TO | TRUSTED | TYPE | UNEQUAL | WITH | XOR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv915 * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _endpos_a_, _menhir_s, (a : 'tv_nonempty_list_atom_), _startpos_a_) = _menhir_stack in + let _startpos = _startpos_a_ in + let _endpos = _endpos_a_ in + let _v : 'tv_expression = let _endpos = _endpos_a_ in + let _symbolstartpos = _startpos_a_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 505 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( match a with + | [hd] -> hd + | _ -> mkexp_ ~loc:_sloc (PEapp a) + ) +# 10005 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv916) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv917 * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv918)) : 'freshtv920) + | MenhirState49 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv923 * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position)) * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv921 * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position)) * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_a1_, _menhir_s, (a1 : 'tv_nonempty_list_atom_), _startpos_a1_), _endpos_a2_, _, (a2 : 'tv_nonempty_list_atom_), _startpos_a2_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos_a1_ in + let _endpos = _endpos_a2_ in + let _v : 'tv_expression = let _endpos = _endpos_a2_ in + let _symbolstartpos = _startpos_a1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 509 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( let x = mkexp_ ~loc:_sloc (PEfield (a1, a2)) in + (* print_endline ("PEfield: " ^ (string_of_p_expression x)); *) x + ) +# 10032 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv922)) : 'freshtv924) + | MenhirState51 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv927 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv925 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos_x_, _menhir_s, (x : 'tv_atom), _startpos_x_), _endpos_xs_, _, (xs : 'tv_nonempty_list_atom_), _startpos_xs_) = _menhir_stack in + let _startpos = _startpos_x_ in + let _endpos = _endpos_xs_ in + let _v : 'tv_nonempty_list_atom_ = +# 223 "" + ( x :: xs ) +# 10046 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_nonempty_list_atom_ _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv926)) : 'freshtv928) + | _ -> + _menhir_fail () + +and _menhir_reduce166 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _v : 'tv_method_kind = +# 378 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( MKnormal ) +# 10057 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_kind _menhir_env _menhir_stack _menhir_s _v + +and _menhir_run274 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv911) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + ((let _1 = () in + let _v : 'tv_method_kind = +# 384 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( MKrefined ) +# 10071 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_kind _menhir_env _menhir_stack _menhir_s _v) : 'freshtv912) + +and _menhir_run276 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv909) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _1 = () in + let _v : 'tv_method_kind = +# 383 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( MKlogical ) +# 10087 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_kind _menhir_env _menhir_stack _menhir_s _v) : 'freshtv910) + +and _menhir_run277 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | CONST -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv903 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv901 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _2 = () in + let _1 = () in + let _v : 'tv_method_kind = +# 382 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( MKconstghost ) +# 10109 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_kind _menhir_env _menhir_stack _menhir_s _v) : 'freshtv902)) : 'freshtv904) + | IDENT _ -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv905 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _1 = () in + let _v : 'tv_method_kind = +# 380 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( MKghost ) +# 10120 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_kind _menhir_env _menhir_stack _menhir_s _v) : 'freshtv906) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv907 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv908) + +and _menhir_run279 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv897 * _menhir_state) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState280 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState280 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState280 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState280 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState280 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState280 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState280) : 'freshtv898) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv899 * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv900) + +and _menhir_run284 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | GHOST -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv891 * _menhir_state) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv889 * _menhir_state) = Obj.magic _menhir_stack in + let (_startpos__2_ : Lexing.position) = _startpos in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + let _2 = () in + let _1 = () in + let _v : 'tv_method_kind = +# 381 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( MKconstghost ) +# 10187 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_kind _menhir_env _menhir_stack _menhir_s _v) : 'freshtv890)) : 'freshtv892) + | IDENT _ -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv893 * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + let _1 = () in + let _v : 'tv_method_kind = +# 379 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( MKconst ) +# 10198 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_method_kind _menhir_env _menhir_stack _menhir_s _v) : 'freshtv894) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv895 * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv896) + +and _menhir_goto_object_signature_expr : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_object_signature_expr -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + match _menhir_s with + | MenhirState272 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv841 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10218 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | MINUS -> + _menhir_run316 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | WITH -> + _menhir_run300 _menhir_env (Obj.magic _menhir_stack) + | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv837 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10232 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _startpos__1_), _endpos__2_, (_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10237 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__2_), _endpos__4_, _, (_4 : 'tv_object_signature_expr), _startpos__4_) = _menhir_stack in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__4_ in + let _v : 'tv_object_signature_declaration = +# 350 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2, _4 ) +# 10246 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_declaration _menhir_env _menhir_stack _endpos _v _startpos) : 'freshtv838) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv839 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10256 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv840)) : 'freshtv842) + | MenhirState329 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv857 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10265 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | MINUS -> + _menhir_run316 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | WITH -> + _menhir_run300 _menhir_env (Obj.magic _menhir_stack) + | RBRACE | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv853 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10279 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10284 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_), _endpos__3_, _, (_3 : 'tv_object_signature_expr), _startpos__3_) = _menhir_stack in + let _2 = () in + let _v : 'tv_layer_signature_field = +# 408 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _1, _3 ) +# 10290 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv851) = _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_layer_signature_field) = _v in + ((match _menhir_s with + | MenhirState332 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv845 * _menhir_state * 'tv_layer_signature_fields) * _menhir_state) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_layer_signature_field) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv843 * _menhir_state * 'tv_layer_signature_fields) * _menhir_state) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_layer_signature_field) : 'tv_layer_signature_field) = _v in + ((let ((_menhir_stack, _menhir_s, (_1 : 'tv_layer_signature_fields)), _) = _menhir_stack in + let _2 = () in + let _v : 'tv_layer_signature_fields = +# 405 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 10311 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_signature_fields _menhir_env _menhir_stack _menhir_s _v) : 'freshtv844)) : 'freshtv846) + | MenhirState326 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv849) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_layer_signature_field) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv847) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_layer_signature_field) : 'tv_layer_signature_field) = _v in + ((let _v : 'tv_layer_signature_fields = +# 404 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 10326 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_signature_fields _menhir_env _menhir_stack _menhir_s _v) : 'freshtv848)) : 'freshtv850) + | _ -> + _menhir_fail ()) : 'freshtv852)) : 'freshtv854) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv855 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10338 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv856)) : 'freshtv858) + | MenhirState452 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv863 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10347 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | MINUS -> + _menhir_run316 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | WITH -> + _menhir_run300 _menhir_env (Obj.magic _menhir_stack) + | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv859 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10361 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((let (((((_menhir_stack, _endpos__1_, (_1 : 'tv_opt_logical_or_trusted), _startpos__1_), _startpos__2_), _startpos__3_), _endpos__4_, (_4 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10366 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__4_), _endpos__6_, _, (_6 : 'tv_object_signature_expr), _startpos__6_) = _menhir_stack in + let _5 = () in + let _3 = () in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__6_ in + let _v : 'tv_object_signature_declaration = +# 353 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( if _1 <> POnormal then + print_endline "Warning: object signatures should not be marked logical or trusted"; + _4, _6 ) +# 10378 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_declaration _menhir_env _menhir_stack _endpos _v _startpos) : 'freshtv860) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv861 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10388 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv862)) : 'freshtv864) + | MenhirState458 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv879 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10397 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | MINUS -> + _menhir_run316 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | WITH -> + _menhir_run300 _menhir_env (Obj.magic _menhir_stack) + | COMMA | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv875 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10411 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10416 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_), _endpos__3_, _, (_3 : 'tv_object_signature_expr), _startpos__3_) = _menhir_stack in + let _2 = () in + let _v : 'tv_base_slot = +# 671 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _1, _3 ) +# 10422 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv873) = _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_base_slot) = _v in + ((match _menhir_s with + | MenhirState464 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv867 * _menhir_state * 'tv_base_slots)) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_base_slot) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv865 * _menhir_state * 'tv_base_slots)) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_base_slot) : 'tv_base_slot) = _v in + ((let (_menhir_stack, _menhir_s, (_1 : 'tv_base_slots)) = _menhir_stack in + let _2 = () in + let _v : 'tv_base_slots = +# 668 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 10443 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_base_slots _menhir_env _menhir_stack _menhir_s _v) : 'freshtv866)) : 'freshtv868) + | MenhirState455 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv871) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_base_slot) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv869) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_base_slot) : 'tv_base_slot) = _v in + ((let _v : 'tv_base_slots = +# 667 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 10458 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_base_slots _menhir_env _menhir_stack _menhir_s _v) : 'freshtv870)) : 'freshtv872) + | _ -> + _menhir_fail ()) : 'freshtv874)) : 'freshtv876) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv877 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10470 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv878)) : 'freshtv880) + | MenhirState471 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv887 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10479 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv881 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10489 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | CLONE -> + _menhir_run357 _menhir_env (Obj.magic _menhir_stack) MenhirState507 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run356 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState507 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run355 _menhir_env (Obj.magic _menhir_stack) MenhirState507 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState507) : 'freshtv882) + | LBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv883 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10509 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LET -> + _menhir_run474 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState473 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState473) : 'freshtv884) + | MINUS -> + _menhir_run316 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | WITH -> + _menhir_run300 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv885 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10533 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv886)) : 'freshtv888) + | _ -> + _menhir_fail () + +and _menhir_goto_base_layer_signature : _menhir_env -> 'ttv_tail -> 'tv_base_layer_signature -> 'ttv_return = + fun _menhir_env _menhir_stack _v -> + let _menhir_stack = (_menhir_stack, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv835 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10547 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv831 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10557 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run298 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState471 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run273 _menhir_env (Obj.magic _menhir_stack) MenhirState471 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState471) : 'freshtv832) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv833 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10577 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv834)) : 'freshtv836) + +and _menhir_run458 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10584 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run298 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState458 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run273 _menhir_env (Obj.magic _menhir_stack) MenhirState458 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState458 + +and _menhir_run355 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | CLONE -> + _menhir_run357 _menhir_env (Obj.magic _menhir_stack) MenhirState355 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run356 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState355 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run355 _menhir_env (Obj.magic _menhir_stack) MenhirState355 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState355 + +and _menhir_run356 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10619 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv829) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10630 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10634 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_object_expression = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 732 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkobj ~loc:_sloc (POname _1) ) +# 10645 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv830) + +and _menhir_run357 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv825 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10662 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv823 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let ((_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10672 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10676 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__2_ : Lexing.position) = _startpos in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__2_ in + let _v : 'tv_object_expression = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 733 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (mkobj ~loc:_sloc (POclone _2)) +# 10689 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv824)) : 'freshtv826) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv827 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv828) + +and _menhir_goto_layer_signature : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_layer_signature -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v) in + match _menhir_s with + | MenhirState325 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv795 * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10709 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv793 * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10715 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10720 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_), _endpos__5_, _, (_5 : 'tv_layer_signature)) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__5_ in + let _v : 'tv_layer_signature_declaration = +# 397 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3, _5 ) +# 10730 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv791) = _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_layer_signature_declaration) = _v in + let (_startpos : Lexing.position) = _startpos in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv789) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_layer_signature_declaration) = _v in + let (_startpos : Lexing.position) = _startpos in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv787) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let ((_1 : 'tv_layer_signature_declaration) : 'tv_layer_signature_declaration) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _endpos = _endpos__1_ in + let _v : 'tv_declaration = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 226 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( fst _1, mkdecl ~loc:_sloc (PDlayer_sig (snd _1)) ) +# 10754 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv788)) : 'freshtv790)) : 'freshtv792)) : 'freshtv794)) : 'freshtv796) + | MenhirState343 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv799 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv797 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _menhir_s, _startpos__1_), _endpos__2_, _), _endpos__3_, _, (_3 : 'tv_layer_signature)) = _menhir_stack in + let _2 = () in + let _1 = () in + let _endpos = _endpos__3_ in + let _v : 'tv_layer_type = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 416 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( { pLayerBase = mklayersign ~loc:_sloc (PLSconstr []); pLayerSignature = _3; pLayerLoc=(make_loc _sloc) } ) +# 10772 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_type _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv798)) : 'freshtv800) + | MenhirState342 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv805 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv801 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_stack = (_menhir_stack, _endpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run337 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState346 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run326 _menhir_env (Obj.magic _menhir_stack) MenhirState346 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState346) : 'freshtv802) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv803 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv804)) : 'freshtv806) + | MenhirState346 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv809 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv807 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_layer_signature)), _endpos__3_), _endpos__4_, _, (_4 : 'tv_layer_signature)) = _menhir_stack in + let _3 = () in + let _1 = () in + let _endpos = _endpos__4_ in + let _v : 'tv_layer_type = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 414 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( { pLayerBase = _2; pLayerSignature = _4; pLayerLoc =(make_loc _sloc) } ) +# 10819 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_type _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv808)) : 'freshtv810) + | MenhirState361 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv813 * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv811 * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_object_expression), _startpos__1_), _endpos__3_, _, (_3 : 'tv_layer_signature)) = _menhir_stack in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_object_expression = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 734 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkobj ~loc:_sloc (POrelax (_1, _3)) ) +# 10837 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv812)) : 'freshtv814) + | MenhirState455 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv821 * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv817 * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv815 * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + ((let ((_menhir_stack, _startpos__1_), _endpos__2_, _, (_2 : 'tv_layer_signature)) = _menhir_stack in + let _3 = () in + let _1 = () in + let _v : 'tv_base_layer_signature = +# 663 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 ) +# 10860 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_base_layer_signature _menhir_env _menhir_stack _v) : 'freshtv816)) : 'freshtv818) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv819 * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv820)) : 'freshtv822) + | _ -> + _menhir_fail () + +and _menhir_run328 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10876 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv783 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10888 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run298 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState329 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run273 _menhir_env (Obj.magic _menhir_stack) MenhirState329 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState329) : 'freshtv784) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv785 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10908 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv786) + +and _menhir_reduce142 : _menhir_env -> 'ttv_tail * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10916 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) * Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack -> + let (_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 10922 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_) = _menhir_stack in + let _endpos = _endpos__1_ in + let _v : 'tv_layer_signature = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 400 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayersign ~loc:_sloc (PLSname _1) ) +# 10931 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_signature _menhir_env _menhir_stack _endpos _menhir_s _v + +and _menhir_run7 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 10938 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv781) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 10949 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 10953 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _v : 'tv_annotation = let _endpos = _endpos__1_ in + let _startpos = _startpos__1_ in + let _loc = (_startpos, _endpos) in + +# 800 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PAexpr (mkcmd ~loc:_loc (PCyield (mkexp_ ~loc:_loc (PEconstant (CONuint (int_of_string(_1))))))) ) +# 10962 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation _menhir_env _menhir_stack _menhir_s _v) : 'freshtv782) + +and _menhir_run8 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 10969 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run28 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState8 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run22 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState8 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run21 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState8 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run11 _menhir_env (Obj.magic _menhir_stack) MenhirState8 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | STRING _v -> + _menhir_run10 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState8 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run9 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState8 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | COMMA | RBRACKET | RPAREN -> + _menhir_reduce10 _menhir_env (Obj.magic _menhir_stack) MenhirState8 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState8 + +and _menhir_goto_annotations : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_annotations -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v) in + match _menhir_s with + | MenhirState125 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv713 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv709 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11012 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSIGN -> + _menhir_run129 _menhir_env (Obj.magic _menhir_stack) MenhirState127 + | EQUAL -> + _menhir_run128 _menhir_env (Obj.magic _menhir_stack) MenhirState127 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState127) : 'freshtv710) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv711 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv712)) : 'freshtv714) + | MenhirState135 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv719 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv715 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11047 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSIGN -> + _menhir_run129 _menhir_env (Obj.magic _menhir_stack) MenhirState137 + | EQUAL -> + _menhir_run128 _menhir_env (Obj.magic _menhir_stack) MenhirState137 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState137) : 'freshtv716) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv717 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv718)) : 'freshtv720) + | MenhirState147 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv725 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv721 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11082 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSIGN -> + _menhir_run129 _menhir_env (Obj.magic _menhir_stack) MenhirState149 + | EQUAL -> + _menhir_run128 _menhir_env (Obj.magic _menhir_stack) MenhirState149 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState149) : 'freshtv722) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv723 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv724)) : 'freshtv726) + | MenhirState4 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv733 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11109 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv729 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11119 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BAR -> + _menhir_run110 _menhir_env (Obj.magic _menhir_stack) MenhirState242 + | LBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv727) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState242 in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run244 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState243 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState243) : 'freshtv728) + | IDENT _ -> + _menhir_reduce208 _menhir_env (Obj.magic _menhir_stack) MenhirState242 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState242) : 'freshtv730) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv731 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11154 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv732)) : 'freshtv734) + | MenhirState256 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv747 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11163 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * 'tv_constructor_params) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv745 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11169 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * 'tv_constructor_params) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11174 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_), _endpos__2_, (_2 : 'tv_constructor_params)), _endpos__3_, _, (_3 : 'tv_annotations)) = _menhir_stack in + let _endpos = _endpos__3_ in + let _v : 'tv_constructor_declaration = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 297 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( { pTypeConstrName = _1; + pTypeConstrArgs = _2; + pTypeConstrAnnotations = _3; + pTypeConstrLoc = (make_loc _sloc ) } ) +# 11186 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv743) = _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_constructor_declaration) = _v in + ((match _menhir_s with + | MenhirState266 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv737 * Lexing.position * _menhir_state * 'tv_constructor_declarations)) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_constructor_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv735 * Lexing.position * _menhir_state * 'tv_constructor_declarations)) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_constructor_declaration) : 'tv_constructor_declaration) = _v in + ((let (_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_constructor_declarations)) = _menhir_stack in + let _2 = () in + let _endpos = _endpos__3_ in + let _v : 'tv_constructor_declarations = +# 293 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 11211 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_constructor_declarations _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv736)) : 'freshtv738) + | MenhirState254 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv741) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_constructor_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv739) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_constructor_declaration) : 'tv_constructor_declaration) = _v in + ((let _endpos = _endpos__1_ in + let _v : 'tv_constructor_declarations = +# 292 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 11229 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_constructor_declarations _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv740)) : 'freshtv742) + | _ -> + _menhir_fail ()) : 'freshtv744)) : 'freshtv746)) : 'freshtv748) + | MenhirState339 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv755 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11239 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv749) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + _menhir_run342 _menhir_env (Obj.magic _menhir_stack) MenhirState341 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState341) : 'freshtv750) + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv751) = Obj.magic _menhir_stack in + ((let _v : 'tv_layer_signature_annotation = +# 752 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( None ) +# 11262 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_signature_annotation _menhir_env _menhir_stack _v) : 'freshtv752) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv753 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11272 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv754)) : 'freshtv756) + | MenhirState401 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv757 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11281 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSIGN -> + _menhir_run129 _menhir_env (Obj.magic _menhir_stack) MenhirState402 + | EQUAL -> + _menhir_run128 _menhir_env (Obj.magic _menhir_stack) MenhirState402 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState402) : 'freshtv758) + | MenhirState407 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv763 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11299 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv759 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11309 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState409 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState409 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState409 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState409 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState409 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState409 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState409) : 'freshtv760) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv761 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11337 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv762)) : 'freshtv764) + | MenhirState399 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv775 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | WITH -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv771 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | STRING _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv767 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11360 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv765 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + let ((_4 : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11370 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11374 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__4_ : Lexing.position) = _startpos in + ((let ((_menhir_stack, _endpos__1_, _startpos__1_), _endpos__2_, _, (_2 : 'tv_annotations)) = _menhir_stack in + let _3 = () in + let _1 = () in + let _endpos = _endpos__4_ in + let _v : 'tv_external_declaration = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 236 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( "", mkdecl ~loc:_sloc (PDexternal_with (_4, _2)) ) +# 11387 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_external_declaration _menhir_env _menhir_stack _endpos _v) : 'freshtv766)) : 'freshtv768) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv769 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv770)) : 'freshtv772) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv773 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv774)) : 'freshtv776) + | MenhirState474 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv779 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | CONST -> + _menhir_run284 _menhir_env (Obj.magic _menhir_stack) MenhirState475 + | CONSTRUCTOR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv777 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState475 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run488 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState476 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run477 _menhir_env (Obj.magic _menhir_stack) MenhirState476 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState476) : 'freshtv778) + | GHOST -> + _menhir_run277 _menhir_env (Obj.magic _menhir_stack) MenhirState475 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LOGICAL -> + _menhir_run276 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState475 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | REFINED -> + _menhir_run274 _menhir_env (Obj.magic _menhir_stack) MenhirState475 + | IDENT _ -> + _menhir_reduce166 _menhir_env (Obj.magic _menhir_stack) MenhirState475 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState475) : 'freshtv780) + | _ -> + _menhir_fail () + +and _menhir_run13 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11446 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv707) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11457 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11461 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _v : 'tv_annotation = let _endpos = _endpos__1_ in + let _startpos = _startpos__1_ in + let _loc = (_startpos, _endpos) in + +# 799 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PAexpr (mkcmd ~loc:_loc (PCyield (mkexp_ ~loc:_loc (PEconstant (CONint (int_of_string(_1))))))) ) +# 11470 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation _menhir_env _menhir_stack _menhir_s _v) : 'freshtv708) + +and _menhir_run14 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv703 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | INT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv691 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11495 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv687 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11507 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv685 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11515 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11521 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _v : 'tv_annotation = let _endpos = _endpos__4_ in + let _startpos = _startpos__1_ in + let _loc = (_startpos, _endpos) in + +# 801 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PAexpr (mkcmd ~loc:_loc (PCyield (mkexp_ ~loc:_loc (PEconstant (CONaddress _3))))) ) +# 11532 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation _menhir_env _menhir_stack _menhir_s _v) : 'freshtv686)) : 'freshtv688) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv689 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11542 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _), _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv690)) : 'freshtv692) + | UINT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv699 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11553 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv695 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11565 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv693 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11573 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11579 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _v : 'tv_annotation = let _endpos = _endpos__4_ in + let _startpos = _startpos__1_ in + let _loc = (_startpos, _endpos) in + +# 802 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( PAexpr (mkcmd ~loc:_loc (PCyield (mkexp_ ~loc:_loc (PEconstant (CONaddress _3))))) ) +# 11590 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotation _menhir_env _menhir_stack _menhir_s _v) : 'freshtv694)) : 'freshtv696) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv697 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 11600 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _), _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv698)) : 'freshtv700) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv701 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv702)) : 'freshtv704) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv705 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv706) + +and _menhir_run20 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11622 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run28 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState20 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run22 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState20 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run21 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState20 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run11 _menhir_env (Obj.magic _menhir_stack) MenhirState20 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | STRING _v -> + _menhir_run10 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState20 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run9 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState20 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | COMMA | RBRACKET | RPAREN -> + _menhir_reduce10 _menhir_env (Obj.magic _menhir_stack) MenhirState20 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState20 + +and _menhir_run30 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + _menhir_run160 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BEGIN -> + _menhir_run159 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | DENY -> + _menhir_run158 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EMIT -> + _menhir_run156 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FAIL -> + _menhir_run155 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FIRST -> + _menhir_run147 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOLD -> + _menhir_run135 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | FOR -> + _menhir_run125 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | GHOST -> + _menhir_run121 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState30 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IF -> + _menhir_run118 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState30 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + _menhir_run115 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MATCH -> + _menhir_run107 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState30 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState30 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState30 + +and _menhir_goto_type_expression : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_type_expression -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + match _menhir_s with + | MenhirState230 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv591 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv589 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (((((_menhir_stack, _menhir_s, _startpos__1_), _startpos__2_), _endpos_e_, _, (e : 'tv_expression), _startpos_e_), _endpos__4_), _endpos__5_, _, (_5 : 'tv_type_expression), _startpos__5_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__5_ in + let _v : 'tv_type_expression = let _endpos = _endpos__5_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 271 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkfotyp ~loc:_sloc (PTarray (int_of_p_expression e.p_expression_desc, _5)) ) +# 11722 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv590)) : 'freshtv592) + | MenhirState224 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv595 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState232 + | ARROW | ASSIGN | COMMA | CONST | EOF | EQUAL | EVENT | EXTERNAL | INDEXED | LAYER | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv593 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_type_expression), _startpos__2_) = _menhir_stack in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__2_ in + let _v : 'tv_type_expression = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 275 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkfotyp ~loc:_sloc (PTlist _2) ) +# 11746 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv594) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState232) : 'freshtv596) + | MenhirState233 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv599 * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv597 * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _endpos__1_, _menhir_s, (_1 : 'tv_type_expression), _startpos__1_), _), _endpos__3_, _, (_3 : 'tv_type_expression), _startpos__3_) = _menhir_stack in + let _2 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_type_expression = let _endpos = _endpos__3_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 269 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkfotyp ~loc:_sloc (PTprod (_1, _3)) ) +# 11768 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv598)) : 'freshtv600) + | MenhirState223 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv605 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv603 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState235 in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv601 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__3_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + ((let ((_menhir_stack, _menhir_s, _startpos__1_), _endpos__2_, _, (_2 : 'tv_type_expression), _startpos__2_) = _menhir_stack in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__3_ in + let _v : 'tv_type_expression = +# 267 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 ) +# 11795 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv602)) : 'freshtv604) + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState235 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState235) : 'freshtv606) + | MenhirState222 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv609 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv607 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState237 in + ((let _menhir_stack = (_menhir_stack, _endpos, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState238 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState238 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState238 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState238 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState238 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState238 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState238) : 'freshtv608) + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState237 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState237) : 'freshtv610) + | MenhirState238 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv613 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * Lexing.position * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState239 + | ARROW | ASSIGN | COMMA | CONST | EOF | EQUAL | EVENT | EXTERNAL | INDEXED | LAYER | LOGICAL | OBJECT | RBRACE | RBRACKET | RPAREN | SEMICOLON | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv611 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * Lexing.position * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (((((_menhir_stack, _menhir_s, _startpos__1_), _startpos__2_), _endpos__3_, _, (_3 : 'tv_type_expression), _startpos__3_), _endpos__4_, _), _endpos__5_, _, (_5 : 'tv_type_expression), _startpos__5_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__5_ in + let _v : 'tv_type_expression = let _endpos = _endpos__5_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 273 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkfotyp ~loc:_sloc (PTmapping (_3, _5)) ) +# 11864 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv612) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState239) : 'freshtv614) + | MenhirState220 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv617 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11876 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState240 + | CONST | EOF | EVENT | EXTERNAL | LAYER | LOGICAL | OBJECT | SIGNATURE | TRUSTED | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv615 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11888 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _startpos__1_), _endpos__2_, (_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11893 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__2_), _), _endpos_te_, _, (te : 'tv_type_expression), _startpos_te_) = _menhir_stack in + let _3 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos_te_ in + let _v : 'tv_type_declaration = let _endpos = _endpos_te_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 253 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2, mkfotyp ~loc:_sloc te.p_type_FO_desc ) +# 11905 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_declaration _menhir_env _menhir_stack _endpos _v _startpos) : 'freshtv616) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState240) : 'freshtv618) + | MenhirState245 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv631 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11917 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState246 + | RBRACE | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv629 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11929 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _endpos__1_, _menhir_s, (_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 11934 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__1_), _endpos__3_, _, (_3 : 'tv_type_expression), _startpos__3_) = _menhir_stack in + let _2 = () in + let _v : 'tv_field_declaration = +# 288 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _1, _3 ) +# 11940 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv627) = _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_field_declaration) = _v in + ((match _menhir_s with + | MenhirState248 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv621 * _menhir_state * 'tv_field_declarations) * _menhir_state) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_field_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv619 * _menhir_state * 'tv_field_declarations) * _menhir_state) = Obj.magic _menhir_stack in + let (_ : _menhir_state) = _menhir_s in + let ((_3 : 'tv_field_declaration) : 'tv_field_declaration) = _v in + ((let ((_menhir_stack, _menhir_s, (_1 : 'tv_field_declarations)), _) = _menhir_stack in + let _2 = () in + let _v : 'tv_field_declarations = +# 285 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _3 :: _1 ) +# 11961 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_field_declarations _menhir_env _menhir_stack _menhir_s _v) : 'freshtv620)) : 'freshtv622) + | MenhirState243 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv625) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_v : 'tv_field_declaration) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv623) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : 'tv_field_declaration) : 'tv_field_declaration) = _v in + ((let _v : 'tv_field_declarations = +# 284 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [_1] ) +# 11976 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_field_declarations _menhir_env _menhir_stack _menhir_s _v) : 'freshtv624)) : 'freshtv626) + | _ -> + _menhir_fail ()) : 'freshtv628)) : 'freshtv630) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState246) : 'freshtv632) + | MenhirState261 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv643 * Lexing.position) * _menhir_state * 'tv_idents)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv641 * Lexing.position) * _menhir_state * 'tv_idents)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState262 in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv639 * Lexing.position) * _menhir_state * 'tv_idents)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__5_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + ((let (((_menhir_stack, _startpos__1_), _, (_2 : 'tv_idents)), _endpos__4_, _, (_4 : 'tv_type_expression), _startpos__4_) = _menhir_stack in + let _5 = () in + let _3 = () in + let _1 = () in + let _endpos = _endpos__5_ in + let _v : 'tv_constructor_parameter = +# 312 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( List.map (fun i -> i, _4) _2 ) +# 12009 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv637) = _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_constructor_parameter) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv635 * Lexing.position * 'tv_constructor_parameters) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _endpos in + let (_v : 'tv_constructor_parameter) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv633 * Lexing.position * 'tv_constructor_parameters) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let ((_2 : 'tv_constructor_parameter) : 'tv_constructor_parameter) = _v in + ((let (_menhir_stack, _endpos__1_, (_1 : 'tv_constructor_parameters)) = _menhir_stack in + let _endpos = _endpos__2_ in + let _v : 'tv_constructor_parameters = +# 308 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2 :: _1 ) +# 12028 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_constructor_parameters _menhir_env _menhir_stack _endpos _v) : 'freshtv634)) : 'freshtv636)) : 'freshtv638)) : 'freshtv640)) : 'freshtv642) + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState262 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState262) : 'freshtv644) + | MenhirState280 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv647 * _menhir_state)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ARROW -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv645 * _menhir_state)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState281 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState282 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState282 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState282 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState282 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState282 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState282 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState282) : 'freshtv646) + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState281 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState281) : 'freshtv648) + | MenhirState282 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv651 * _menhir_state)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState283 + | RBRACE | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv649 * _menhir_state)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _menhir_s), _endpos__3_, _, (_3 : 'tv_type_expression), _startpos__3_), _), _endpos__5_, _, (_5 : 'tv_type_expression), _startpos__5_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _v : 'tv_object_signature_field = +# 375 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( "constructor", _3, _5, MKconstructor ) +# 12091 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_field _menhir_env _menhir_stack _menhir_s _v) : 'freshtv650) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState283) : 'freshtv652) + | MenhirState291 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv655 * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12103 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ARROW -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv653 * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12113 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState292 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState293 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState293 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState293 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState293 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState293 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState293 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState293) : 'freshtv654) + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState292 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState292) : 'freshtv656) + | MenhirState293 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv659 * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12147 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState294 + | RBRACE | SEMICOLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv657 * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12159 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (((((_menhir_stack, _menhir_s, (_1 : 'tv_method_kind)), _endpos__2_, (_2 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12164 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__2_), _endpos__4_, _, (_4 : 'tv_type_expression), _startpos__4_), _), _endpos__6_, _, (_6 : 'tv_type_expression), _startpos__6_) = _menhir_stack in + let _5 = () in + let _3 = () in + let _v : 'tv_object_signature_field = +# 373 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( _2, _4, _6, _1 ) +# 12171 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_field _menhir_env _menhir_stack _menhir_s _v) : 'freshtv658) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState294) : 'freshtv660) + | MenhirState409 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv663 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12183 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ARROW -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv661 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12193 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState410 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState411 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState411 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState411 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState411 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState411 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState411 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState411) : 'freshtv662) + | ASSIGN -> + _menhir_run129 _menhir_env (Obj.magic _menhir_stack) MenhirState410 + | EQUAL -> + _menhir_run128 _menhir_env (Obj.magic _menhir_stack) MenhirState410 + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState410 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState410) : 'freshtv664) + | MenhirState411 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv665 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12231 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSIGN -> + _menhir_run129 _menhir_env (Obj.magic _menhir_stack) MenhirState412 + | EQUAL -> + _menhir_run128 _menhir_env (Obj.magic _menhir_stack) MenhirState412 + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState412 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState412) : 'freshtv666) + | MenhirState419 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv667 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12251 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSIGN -> + _menhir_run129 _menhir_env (Obj.magic _menhir_stack) MenhirState420 + | EQUAL -> + _menhir_run128 _menhir_env (Obj.magic _menhir_stack) MenhirState420 + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState420 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState420) : 'freshtv668) + | MenhirState432 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv675 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12271 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | INDEXED -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv671) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState433 in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv669) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + ((let _1 = () in + let _v : 'tv_indexed_opt = +# 339 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( true ) +# 12288 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_indexed_opt _menhir_env _menhir_stack _menhir_s _v) : 'freshtv670)) : 'freshtv672) + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState433 + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv673) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState433 in + ((let _v : 'tv_indexed_opt = +# 338 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( false ) +# 12300 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_indexed_opt _menhir_env _menhir_stack _menhir_s _v) : 'freshtv674) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState433) : 'freshtv676) + | MenhirState480 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv679 * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState481 + | COMMA | EQUAL | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv677 * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s), _endpos__2_, _, (_2 : 'tv_type_expression), _startpos__2_) = _menhir_stack in + let _1 = () in + let _v : 'tv_opt_type_annotation = +# 729 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( Some _2 ) +# 12323 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_opt_type_annotation _menhir_env _menhir_stack _menhir_s _v) : 'freshtv678) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState481) : 'freshtv680) + | MenhirState495 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv683 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12335 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSIGN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv681 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12345 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState496 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState497 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState497 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState497 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState497 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState497 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState497 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState497 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState497 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState497 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState497) : 'freshtv682) + | STAR -> + _menhir_run233 _menhir_env (Obj.magic _menhir_stack) MenhirState496 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState496) : 'freshtv684) + | _ -> + _menhir_fail () + +and _menhir_goto_opt_bar : _menhir_env -> 'ttv_tail -> _menhir_state -> 'tv_opt_bar -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _v -> + let _menhir_stack = (_menhir_stack, _menhir_s, _v) in + match _menhir_s with + | MenhirState109 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv583 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * _menhir_state * 'tv_opt_bar) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run193 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState111 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACKET -> + _menhir_run112 _menhir_env (Obj.magic _menhir_stack) MenhirState111 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState111) : 'freshtv584) + | MenhirState242 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv585 * _menhir_state * 'tv_opt_bar) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run255 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState254 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState254) : 'freshtv586) + | MenhirState426 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv587) * _menhir_state * 'tv_opt_bar) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run428 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState427 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState427) : 'freshtv588) + | _ -> + _menhir_fail () + +and _menhir_run36 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12431 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv579 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12443 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState37 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState37 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState37 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState37 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState37 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState37 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState37 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState37 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState37 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState37) : 'freshtv580) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv581 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12477 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv582) + +and _menhir_goto_atom : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> 'tv_atom -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv577 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState51 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState51 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState51 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv573 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState51 in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState52 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState52 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState52 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState52 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState52 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState52 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState52 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState52 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState52 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState52) : 'freshtv574) + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState51 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState51 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | ASSIGN | BAR | BARBAR | BITAND | COMMA | CONJUNCTION | CONST | DISJUNCTION | DO | DOT | ELSE | END | EOF | EQUAL | EVENT | EXTERNAL | GREATER | GREATEREQ | IN | LAYER | LESS | LESSEQ | LET | LOGICAL | MINUS | MOD | OBJECT | PLUS | RBRACE | RBRACKET | RPAREN | SEMICOLON | SHL | SHR | SIGNATURE | SLASH | STAR | THEN | TO | TRUSTED | TYPE | UNEQUAL | WITH | XOR -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv575 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _endpos_x_, _menhir_s, (x : 'tv_atom), _startpos_x_) = _menhir_stack in + let _startpos = _startpos_x_ in + let _endpos = _endpos_x_ in + let _v : 'tv_nonempty_list_atom_ = +# 221 "" + ( [ x ] ) +# 12540 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_nonempty_list_atom_ _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv576) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState51) : 'freshtv578) + +and _menhir_run273 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | CONST -> + _menhir_run284 _menhir_env (Obj.magic _menhir_stack) MenhirState273 + | CONSTRUCTOR -> + _menhir_run279 _menhir_env (Obj.magic _menhir_stack) MenhirState273 + | GHOST -> + _menhir_run277 _menhir_env (Obj.magic _menhir_stack) MenhirState273 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LOGICAL -> + _menhir_run276 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState273 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv571 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState273 in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv569 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _2 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__2_ in + let _v : 'tv_object_signature_expr = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 359 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mksig ~loc:_sloc (PSconstr []) ) +# 12583 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_expr _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv570)) : 'freshtv572) + | REFINED -> + _menhir_run274 _menhir_env (Obj.magic _menhir_stack) MenhirState273 + | IDENT _ -> + _menhir_reduce166 _menhir_env (Obj.magic _menhir_stack) MenhirState273 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState273 + +and _menhir_run298 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12598 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv567) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12609 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12613 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_object_signature_expr = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 358 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mksig ~loc:_sloc (PSname _1) ) +# 12624 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_object_signature_expr _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv568) + +and _menhir_goto_opt_logical_or_trusted : _menhir_env -> 'ttv_tail -> Lexing.position -> 'tv_opt_logical_or_trusted -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv565 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | OBJECT -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv561 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv547 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12651 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv531 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12663 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | CLONE -> + _menhir_run357 _menhir_env (Obj.magic _menhir_stack) MenhirState468 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run356 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState468 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run355 _menhir_env (Obj.magic _menhir_stack) MenhirState468 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState468) : 'freshtv532) + | LPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv541) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv535) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState455 in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12694 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + _menhir_run458 _menhir_env (Obj.magic _menhir_stack) + | RPAREN -> + _menhir_reduce142 _menhir_env (Obj.magic _menhir_stack) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv533 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12712 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv534)) : 'freshtv536) + | LBRACE -> + _menhir_run326 _menhir_env (Obj.magic _menhir_stack) MenhirState455 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv539 * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState455 in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv537 * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + ((let (_menhir_stack, _startpos__1_) = _menhir_stack in + let _2 = () in + let _1 = () in + let _v : 'tv_base_layer_signature = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 662 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayersign ~loc:_sloc (PLSconstr []) ) +# 12737 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_base_layer_signature _menhir_env _menhir_stack _v) : 'freshtv538)) : 'freshtv540) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState455) : 'freshtv542) + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv543) = Obj.magic _menhir_stack in + ((let (_, _endpos__0_) = Obj.magic _menhir_stack in + let _v : 'tv_base_layer_signature = let _endpos = _endpos__0_ in + let _symbolstartpos = _endpos in + let _sloc = (_symbolstartpos, _endpos) in + +# 661 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayersign ~loc:_sloc (PLSconstr []) ) +# 12754 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_base_layer_signature _menhir_env _menhir_stack _v) : 'freshtv544) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv545 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12764 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv546)) : 'freshtv548) + | SIGNATURE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv557 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv553 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12782 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv549 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12794 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run298 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState452 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run273 _menhir_env (Obj.magic _menhir_stack) MenhirState452 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState452) : 'freshtv550) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv551 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12814 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv552)) : 'freshtv554) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv555 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv556)) : 'freshtv558) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv559 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv560)) : 'freshtv562) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv563 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv564)) : 'freshtv566) + +and _menhir_run326 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run328 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState326 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RBRACE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv529 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState326 in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv527 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _2 = () in + let _1 = () in + let _endpos = _endpos__2_ in + let _v : 'tv_layer_signature = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 401 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mklayersign ~loc:_sloc (PLSconstr []) ) +# 12864 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_layer_signature _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv528)) : 'freshtv530) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState326 + +and _menhir_run337 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 12875 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + _menhir_reduce142 _menhir_env (Obj.magic _menhir_stack) + +and _menhir_reduce19 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let (_, _endpos) = Obj.magic _menhir_stack in + let _v : 'tv_annotations = +# 788 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [] ) +# 12888 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotations _menhir_env _menhir_stack _endpos _menhir_s _v + +and _menhir_run5 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv523 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + _menhir_run30 _menhir_env (Obj.magic _menhir_stack) MenhirState6 + | IDENT _v -> + _menhir_run20 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState6 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run14 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState6 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run13 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState6 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv521 * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState6 in + ((let _menhir_stack = (_menhir_stack, _endpos, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv517 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv515 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _menhir_s, _startpos__1_), _startpos__2_), _endpos__3_, _) = _menhir_stack in + let _4 = () in + let _3 = () in + let _2 = () in + let _1 = () in + let _endpos = _endpos__4_ in + let _v : 'tv_annotations = +# 789 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [] ) +# 12940 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_annotations _menhir_env _menhir_stack _endpos _menhir_s _v) : 'freshtv516)) : 'freshtv518) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv519 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv520)) : 'freshtv522) + | STRING _v -> + _menhir_run8 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState6 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run7 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState6 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState6) : 'freshtv524) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv525 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv526) + +and _menhir_run221 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv511 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState222 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState222 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState222 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState222 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState222 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState222 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState222) : 'freshtv512) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv513 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv514) + +and _menhir_run223 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState223 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState223 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState223 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState223 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState223 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState223 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState223 + +and _menhir_run224 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState224 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState224 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState224 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState224 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState224 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState224 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState224 + +and _menhir_run225 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv509) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_type_expression = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 265 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkfotyp ~loc:_sloc (PTbuiltin Taddress) ) +# 13067 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv510) + +and _menhir_run226 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13074 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv507) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13085 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13089 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_type_expression = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 260 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( let x = try Hashtbl.find builtin_type_table _1 + with Not_found -> PTname _1 + in + mkfotyp ~loc:_sloc (x) + ) +# 13104 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_type_expression _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv508) + +and _menhir_run227 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv503 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState228 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState228 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState228 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState228 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState228 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState228 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState228 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState228 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState228 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState228) : 'freshtv504) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv505 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv506) + +and _menhir_reduce208 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _v : 'tv_opt_bar = +# 833 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( () ) +# 13157 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_opt_bar _menhir_env _menhir_stack _menhir_s _v + +and _menhir_run110 : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv501) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = _menhir_s in + ((let _1 = () in + let _v : 'tv_opt_bar = +# 834 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( () ) +# 13171 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_opt_bar _menhir_env _menhir_stack _menhir_s _v) : 'freshtv502) + +and _menhir_errorcase : _menhir_env -> 'ttv_tail -> _menhir_state -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s -> + match _menhir_s with + | MenhirState507 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv131 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13183 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv132) + | MenhirState503 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv133 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13192 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) * _menhir_state * 'tv_object_fields_and_methods) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv134) + | MenhirState501 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv135 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13201 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv136) + | MenhirState499 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv137 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13210 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_method_param) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv138) + | MenhirState497 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv139 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13219 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv140) + | MenhirState496 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv141 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13228 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv142) + | MenhirState495 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv143 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13237 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv144) + | MenhirState494 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv145 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13246 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv146) + | MenhirState491 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv147 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state) * _menhir_state * 'tv_method_param) * _menhir_state * 'tv_opt_type_annotation)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv148) + | MenhirState489 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv149 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state) * _menhir_state * 'tv_method_param) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv150) + | MenhirState485 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv151 * _menhir_state * 'tv_method_parameters)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv152) + | MenhirState481 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv153 * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv154) + | MenhirState480 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv155 * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv156) + | MenhirState479 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv157 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13280 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv158) + | MenhirState477 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv159 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv160) + | MenhirState476 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv161 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv162) + | MenhirState475 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv163 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv164) + | MenhirState474 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv165 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv166) + | MenhirState473 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv167 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13309 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv168) + | MenhirState471 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv169 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13318 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * 'tv_base_layer_signature)) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv170) + | MenhirState468 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv171 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13326 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv172) + | MenhirState464 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv173 * _menhir_state * 'tv_base_slots)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv174) + | MenhirState458 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv175 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13339 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv176) + | MenhirState455 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv177 * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv178) + | MenhirState452 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv179 * Lexing.position * 'tv_opt_logical_or_trusted * Lexing.position) * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13352 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv180) + | MenhirState445 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv181) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13360 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv182) + | MenhirState439 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv183 * Lexing.position * _menhir_state * 'tv_event_declarations)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv184) + | MenhirState433 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv185 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13373 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv186) + | MenhirState432 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv187 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13382 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv188) + | MenhirState427 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv189) * _menhir_state * 'tv_opt_bar) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv190) + | MenhirState426 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv191) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv192) + | MenhirState420 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv193 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13399 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv194) + | MenhirState419 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv195 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13408 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv196) + | MenhirState412 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv197 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13417 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv198) + | MenhirState411 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((('freshtv199 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13426 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv200) + | MenhirState410 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv201 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13435 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv202) + | MenhirState409 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv203 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13444 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv204) + | MenhirState407 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv205 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13453 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv206) + | MenhirState402 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv207 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13462 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv208) + | MenhirState401 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv209 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13471 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv210) + | MenhirState399 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv211 * Lexing.position * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv212) + | MenhirState396 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv213 * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv214) + | MenhirState393 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv215 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv216) + | MenhirState391 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv217 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv218) + | MenhirState387 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv219 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv220) + | MenhirState385 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv221 * Lexing.position * _menhir_state * 'tv_layer_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv222) + | MenhirState376 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv223 * _menhir_state * 'tv_layer_slots_plus) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv224) + | MenhirState375 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv225 * _menhir_state * 'tv_layer_slots_plus) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv226) + | MenhirState371 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv227 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 13518 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position)) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _, _menhir_s, _), _), _, _, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv228) + | MenhirState367 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv229 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 13527 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position)) = Obj.magic _menhir_stack in + ((let ((((_menhir_stack, _, _menhir_s, _), _), _, _, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv230) + | MenhirState361 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv231 * Lexing.position * _menhir_state * 'tv_object_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv232) + | MenhirState355 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv233 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv234) + | MenhirState354 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv235 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13546 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv236) + | MenhirState352 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv237 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv238) + | MenhirState351 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv239 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv240) + | MenhirState350 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv241 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13565 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * 'tv_layer_signature_annotation)) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv242) + | MenhirState346 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv243 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_layer_signature) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv244) + | MenhirState343 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv245 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv246) + | MenhirState342 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv247 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv248) + | MenhirState341 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv249) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv250) + | MenhirState339 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv251 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13593 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv252) + | MenhirState332 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv253 * _menhir_state * 'tv_layer_signature_fields) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv254) + | MenhirState331 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv255 * _menhir_state * Lexing.position) * _menhir_state * 'tv_layer_signature_fields) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv256) + | MenhirState329 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv257 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13611 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv258) + | MenhirState326 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv259 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv260) + | MenhirState325 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv261 * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13625 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv262) + | MenhirState317 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv263 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position) * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _, _), _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv264) + | MenhirState312 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv265 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _, _), _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv266) + | MenhirState304 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv267 * _menhir_state * 'tv_idents_semi_sep) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv268) + | MenhirState302 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv269 * Lexing.position * _menhir_state * 'tv_object_signature_expr * Lexing.position)) * Lexing.position * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _, _), _, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv270) + | MenhirState294 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv271 * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13653 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv272) + | MenhirState293 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv273 * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13662 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv274) + | MenhirState292 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv275 * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13671 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv276) + | MenhirState291 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv277 * _menhir_state * 'tv_method_kind) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13680 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv278) + | MenhirState287 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv279 * _menhir_state * 'tv_object_signature_fields) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv280) + | MenhirState286 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv281 * _menhir_state * Lexing.position) * _menhir_state * 'tv_object_signature_fields) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv282) + | MenhirState283 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv283 * _menhir_state)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv284) + | MenhirState282 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv285 * _menhir_state)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv286) + | MenhirState281 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv287 * _menhir_state)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv288) + | MenhirState280 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv289 * _menhir_state)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv290) + | MenhirState273 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv291 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv292) + | MenhirState272 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv293 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13724 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv294) + | MenhirState266 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv295 * Lexing.position * _menhir_state * 'tv_constructor_declarations)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv296) + | MenhirState262 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv297 * Lexing.position) * _menhir_state * 'tv_idents)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv298) + | MenhirState261 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv299 * Lexing.position) * _menhir_state * 'tv_idents)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv300) + | MenhirState259 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv301 * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv302) + | MenhirState256 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv303 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13751 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * 'tv_constructor_params) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _, _), _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv304) + | MenhirState254 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv305 * _menhir_state * 'tv_opt_bar) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv306) + | MenhirState248 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv307 * _menhir_state * 'tv_field_declarations) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv308) + | MenhirState247 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv309 * _menhir_state * Lexing.position) * _menhir_state * 'tv_field_declarations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv310) + | MenhirState246 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv311 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13775 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv312) + | MenhirState245 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv313 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13784 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv314) + | MenhirState243 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv315 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv316) + | MenhirState242 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv317 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13798 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv318) + | MenhirState240 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv319 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13807 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv320) + | MenhirState239 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv321 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * Lexing.position * _menhir_state) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv322) + | MenhirState238 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv323 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * Lexing.position * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv324) + | MenhirState237 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv325 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv326) + | MenhirState235 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv327 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv328) + | MenhirState233 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv329 * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv330) + | MenhirState232 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv331 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_type_expression * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv332) + | MenhirState230 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv333 * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv334) + | MenhirState228 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv335 * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv336) + | MenhirState224 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv337 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv338) + | MenhirState223 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv339 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv340) + | MenhirState222 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv341 * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv342) + | MenhirState220 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv343 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13871 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv344) + | MenhirState211 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv345 * _menhir_state * 'tv_annotations_plus)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv346) + | MenhirState205 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv347 * _menhir_state * 'tv_match_clauses)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv348) + | MenhirState201 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv349 * _menhir_state * 'tv_match_pattern)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv350) + | MenhirState193 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv351 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13895 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv352) + | MenhirState190 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv353 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13904 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv354) + | MenhirState187 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv355 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv356) + | MenhirState183 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((('freshtv357 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13918 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) * Lexing.position * _menhir_state * 'tv_commands)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv358) + | MenhirState178 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((((('freshtv359 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13927 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv360) + | MenhirState176 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((((((('freshtv361 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13936 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv362) + | MenhirState174 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((((('freshtv363 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13945 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv364) + | MenhirState170 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv365 * Lexing.position * _menhir_state * 'tv_command * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv366) + | MenhirState162 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv367 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv368) + | MenhirState160 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv369 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv370) + | MenhirState159 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv371 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv372) + | MenhirState158 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv373 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv374) + | MenhirState156 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv375 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv376) + | MenhirState154 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv377 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13984 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv378) + | MenhirState152 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv379 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 13993 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv380) + | MenhirState150 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv381 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14002 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv382) + | MenhirState149 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv383 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14011 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv384) + | MenhirState147 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv385 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv386) + | MenhirState146 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((((((('freshtv387 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14025 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14029 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv388) + | MenhirState144 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((((('freshtv389 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14038 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14042 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv390) + | MenhirState143 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (((((((('freshtv391 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14051 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14055 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv392) + | MenhirState140 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv393 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14064 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv394) + | MenhirState138 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv395 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14073 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv396) + | MenhirState137 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv397 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14082 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv398) + | MenhirState135 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv399 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv400) + | MenhirState134 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((((('freshtv401 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14096 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv402) + | MenhirState132 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((((('freshtv403 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14105 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv404) + | MenhirState130 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv405 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14114 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) * _menhir_state * 'tv_eq_or_assign) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv406) + | MenhirState127 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv407 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_annotations) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14123 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv408) + | MenhirState125 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv409 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv410) + | MenhirState124 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv411 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14137 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv412) + | MenhirState121 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv413 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv414) + | MenhirState120 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv415 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv416) + | MenhirState118 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv417 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv418) + | MenhirState117 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv419 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14161 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv420) + | MenhirState114 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv421 * _menhir_state * Lexing.position) * Lexing.position)) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv422) + | MenhirState111 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ((('freshtv423 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) * _menhir_state * 'tv_opt_bar) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv424) + | MenhirState109 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv425 * _menhir_state * Lexing.position) * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv426) + | MenhirState107 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv427 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv428) + | MenhirState102 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv429 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv430) + | MenhirState96 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv431 * _menhir_state * 'tv_struct_fields) * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv432) + | MenhirState95 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv433 * _menhir_state * Lexing.position) * _menhir_state * 'tv_struct_fields) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv434) + | MenhirState89 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv435 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv436) + | MenhirState87 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv437 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv438) + | MenhirState85 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv439 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv440) + | MenhirState83 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv441 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv442) + | MenhirState81 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv443 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv444) + | MenhirState79 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv445 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv446) + | MenhirState77 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv447 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv448) + | MenhirState74 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv449 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv450) + | MenhirState72 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv451 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv452) + | MenhirState70 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv453 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv454) + | MenhirState68 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv455 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv456) + | MenhirState66 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv457 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv458) + | MenhirState64 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv459 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv460) + | MenhirState62 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv461 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv462) + | MenhirState60 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv463 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv464) + | MenhirState58 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv465 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv466) + | MenhirState56 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv467 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv468) + | MenhirState54 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv469 * Lexing.position * _menhir_state * 'tv_expression * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv470) + | MenhirState52 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv471 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv472) + | MenhirState51 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv473 * Lexing.position * _menhir_state * 'tv_atom * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv474) + | MenhirState49 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv475 * Lexing.position * _menhir_state * 'tv_nonempty_list_atom_ * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv476) + | MenhirState47 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv477 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv478) + | MenhirState46 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv479 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv480) + | MenhirState37 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv481 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14320 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position)) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv482) + | MenhirState35 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv483 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv484) + | MenhirState33 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv485 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv486) + | MenhirState32 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv487 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv488) + | MenhirState30 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv489 * _menhir_state) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv490) + | MenhirState20 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv491 * Lexing.position * _menhir_state * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14349 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv492) + | MenhirState11 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv493 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv494) + | MenhirState8 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv495 * Lexing.position * _menhir_state * ( +# 100 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14363 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv496) + | MenhirState6 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv497 * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv498) + | MenhirState4 -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv499 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14377 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv500) + +and _menhir_run31 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14384 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv129) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14395 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14399 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_atom = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 494 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEconstant (CONuint (int_of_string(_1)))) ) +# 14410 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_atom _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv130) + +and _menhir_run32 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState32 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState32 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState32 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState32 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState32 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState32 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState32 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState32 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState32 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState32 + +and _menhir_run33 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState33 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState33 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState33 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState33 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState33 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState33 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState33 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState33 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv127 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState33 in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv125 * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__2_ : Lexing.position) = _endpos in + let (_ : _menhir_state) = _menhir_s in + ((let (_menhir_stack, _menhir_s, _startpos__1_) = _menhir_stack in + let _2 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__2_ in + let _v : 'tv_atom = let _endpos = _endpos__2_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 497 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEconstant CONunit) ) +# 14486 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_atom _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv126)) : 'freshtv128) + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState33 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState33 + +and _menhir_run35 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run36 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState35 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState35 + +and _menhir_run38 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14512 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv123) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14523 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14527 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_atom = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 493 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEconstant (CONint (int_of_string(_1)))) ) +# 14538 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_atom _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv124) + +and _menhir_run39 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv119 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | INT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv107 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14563 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv103 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14575 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv101 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14583 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14589 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__4_ in + let _v : 'tv_atom = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 495 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEconstant (CONaddress _3)) ) +# 14602 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_atom _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv102)) : 'freshtv104) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv105 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 92 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14612 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _), _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv106)) : 'freshtv108) + | UINT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv115 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14623 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | RPAREN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv111 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14635 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv109 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14643 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos__4_ : Lexing.position) = _endpos in + ((let (((_menhir_stack, _endpos__1_, _menhir_s, _startpos__1_), _startpos__2_), _endpos__3_, (_3 : ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14649 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )), _startpos__3_) = _menhir_stack in + let _4 = () in + let _2 = () in + let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__4_ in + let _v : 'tv_atom = let _endpos = _endpos__4_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 496 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( mkexp_ ~loc:_sloc (PEconstant (CONaddress _3)) ) +# 14662 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_atom _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv110)) : 'freshtv112) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv113 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) * Lexing.position * ( +# 93 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (string) +# 14672 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let (((_menhir_stack, _, _menhir_s, _), _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv114)) : 'freshtv116) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv117 * Lexing.position * _menhir_state * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _, _menhir_s, _), _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv118)) : 'freshtv120) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv121 * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv122) + +and _menhir_run45 : _menhir_env -> 'ttv_tail -> Lexing.position -> _menhir_state -> ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14694 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _menhir_s _v _startpos -> + let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv99) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_menhir_s : _menhir_state) = _menhir_s in + let ((_1 : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14705 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14709 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_atom = let _endpos = _endpos__1_ in + let _symbolstartpos = _startpos__1_ in + let _sloc = (_symbolstartpos, _endpos) in + +# 492 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( let x = try Hashtbl.find constant_table _1 with Not_found -> PEglob _1 in mkexp_ ~loc:_sloc (x) ) +# 14720 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_atom _menhir_env _menhir_stack _endpos _menhir_s _v _startpos) : 'freshtv100) + +and _menhir_run46 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState46 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState46 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState46 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState46 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState46 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState46 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState46 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState46 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState46 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState46 + +and _menhir_run47 : _menhir_env -> 'ttv_tail -> _menhir_state -> Lexing.position -> 'ttv_return = + fun _menhir_env _menhir_stack _menhir_s _startpos -> + let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState47 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState47 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState47 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState47 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState47 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState47 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState47 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState47 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState47 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState47 + +and _menhir_goto_declarations : _menhir_env -> 'ttv_tail -> Lexing.position -> 'tv_declarations -> 'ttv_return = + fun _menhir_env _menhir_stack _endpos _v -> + let _menhir_stack = (_menhir_stack, _endpos, _v) in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv97 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + ((assert (not _menhir_env._menhir_error); + let _tok = _menhir_env._menhir_token in + match _tok with + | CONST -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv13) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv9) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14803 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv5) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14815 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BANG -> + _menhir_run47 _menhir_env (Obj.magic _menhir_stack) MenhirState445 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | BITNOT -> + _menhir_run46 _menhir_env (Obj.magic _menhir_stack) MenhirState445 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run45 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState445 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run39 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState445 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | INT _v -> + _menhir_run38 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState445 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run35 _menhir_env (Obj.magic _menhir_stack) MenhirState445 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run33 _menhir_env (Obj.magic _menhir_stack) MenhirState445 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MINUS -> + _menhir_run32 _menhir_env (Obj.magic _menhir_stack) MenhirState445 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | UINT _v -> + _menhir_run31 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState445 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState445) : 'freshtv6) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv7) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14849 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv8)) : 'freshtv10) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv11) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv12)) : 'freshtv14) + | EOF -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv23 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv21 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _endpos__1_, (_1 : 'tv_declarations)) = _menhir_stack in + let _2 = () in + let _v : ( +# 207 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Parsetree.p_file_structure) +# 14868 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) = +# 213 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( List.rev _1 ) +# 14872 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv19) = _menhir_stack in + let (_v : ( +# 207 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Parsetree.p_file_structure) +# 14879 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv17) = Obj.magic _menhir_stack in + let (_v : ( +# 207 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Parsetree.p_file_structure) +# 14886 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + ((let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv15) = Obj.magic _menhir_stack in + let ((_1 : ( +# 207 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Parsetree.p_file_structure) +# 14893 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) : ( +# 207 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Parsetree.p_file_structure) +# 14897 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + (Obj.magic _1 : 'freshtv16)) : 'freshtv18)) : 'freshtv20)) : 'freshtv22)) : 'freshtv24) + | EVENT -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv25) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | BAR -> + _menhir_run110 _menhir_env (Obj.magic _menhir_stack) MenhirState426 + | IDENT _ -> + _menhir_reduce208 _menhir_env (Obj.magic _menhir_stack) MenhirState426 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState426) : 'freshtv26) + | EXTERNAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv49) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSERT -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv35 * Lexing.position * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState399 in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv31 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14939 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | COLON -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv27 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14951 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState419 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState419 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState419 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState419 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState419 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState419 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState419) : 'freshtv28) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv29 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 14979 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let ((_menhir_stack, _menhir_s, _), _, _, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv30)) : 'freshtv32) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv33 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv34)) : 'freshtv36) + | LBRACKET -> + _menhir_run5 _menhir_env (Obj.magic _menhir_stack) MenhirState399 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LET -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv41 * Lexing.position * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_menhir_s : _menhir_state) = MenhirState399 in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv37 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15009 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + _menhir_run5 _menhir_env (Obj.magic _menhir_stack) MenhirState407 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | COLON -> + _menhir_reduce19 _menhir_env (Obj.magic _menhir_stack) MenhirState407 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState407) : 'freshtv38) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv39 * Lexing.position * Lexing.position) * Lexing.position * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv40)) : 'freshtv42) + | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv47 * Lexing.position * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState399 in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _menhir_s, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv43 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15047 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + _menhir_run5 _menhir_env (Obj.magic _menhir_stack) MenhirState401 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | ASSIGN | EQUAL -> + _menhir_reduce19 _menhir_env (Obj.magic _menhir_stack) MenhirState401 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState401) : 'freshtv44) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv45 * Lexing.position * Lexing.position) * _menhir_state * Lexing.position) = Obj.magic _menhir_stack in + ((let (_menhir_stack, _menhir_s, _) = _menhir_stack in + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) _menhir_s) : 'freshtv46)) : 'freshtv48) + | WITH -> + _menhir_reduce19 _menhir_env (Obj.magic _menhir_stack) MenhirState399 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState399) : 'freshtv50) + | LAYER -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv65) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv51 * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15090 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | LBRACKET -> + _menhir_run5 _menhir_env (Obj.magic _menhir_stack) MenhirState339 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | COLON | EQUAL -> + _menhir_reduce19 _menhir_env (Obj.magic _menhir_stack) MenhirState339 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState339) : 'freshtv52) + | SIGNATURE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv61 * Lexing.position) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv57 * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15120 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv53 * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15132 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run337 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState325 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run326 _menhir_env (Obj.magic _menhir_stack) MenhirState325 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState325) : 'freshtv54) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : (('freshtv55 * Lexing.position) * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15152 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv56)) : 'freshtv58) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv59 * Lexing.position) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv60)) : 'freshtv62) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv63 * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv64)) : 'freshtv66) + | LOGICAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv69) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv67) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_opt_logical_or_trusted = +# 825 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( POlogical ) +# 15183 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_opt_logical_or_trusted _menhir_env _menhir_stack _endpos _v _startpos) : 'freshtv68)) : 'freshtv70) + | SIGNATURE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv79) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv75 * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15201 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | EQUAL -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv71 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15213 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + ((let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + _menhir_run298 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState272 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LBRACE -> + _menhir_run273 _menhir_env (Obj.magic _menhir_stack) MenhirState272 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState272) : 'freshtv72) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv73 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15233 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv74)) : 'freshtv76) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv77 * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv78)) : 'freshtv80) + | TRUSTED -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv83) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv81) = Obj.magic _menhir_stack in + let (_endpos__1_ : Lexing.position) = _endpos in + let (_startpos__1_ : Lexing.position) = _startpos in + ((let _1 = () in + let _startpos = _startpos__1_ in + let _endpos = _endpos__1_ in + let _v : 'tv_opt_logical_or_trusted = +# 826 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( POtrusted ) +# 15258 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_opt_logical_or_trusted _menhir_env _menhir_stack _endpos _v _startpos) : 'freshtv82)) : 'freshtv84) + | TYPE -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv91) = Obj.magic _menhir_stack in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | IDENT _v -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv87 * Lexing.position) = Obj.magic _menhir_stack in + let (_endpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_curr_p in + let (_v : ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15276 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + )) = _v in + let (_startpos : Lexing.position) = _menhir_env._menhir_lexbuf.Lexing.lex_start_p in + ((let _menhir_stack = (_menhir_stack, _endpos, _v, _startpos) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ASSIGN -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : ('freshtv85 * Lexing.position) * Lexing.position * ( +# 91 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Astcommon.ident) +# 15288 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + ) * Lexing.position) = Obj.magic _menhir_stack in + let (_menhir_s : _menhir_state) = MenhirState4 in + ((let _menhir_stack = (_menhir_stack, _menhir_s) in + let _menhir_env = _menhir_discard _menhir_env in + let _tok = _menhir_env._menhir_token in + match _tok with + | ARRAY -> + _menhir_run227 _menhir_env (Obj.magic _menhir_stack) MenhirState220 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENT _v -> + _menhir_run226 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState220 _v _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | IDENTITY -> + _menhir_run225 _menhir_env (Obj.magic _menhir_stack) _menhir_env._menhir_lexbuf.Lexing.lex_curr_p MenhirState220 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LIST -> + _menhir_run224 _menhir_env (Obj.magic _menhir_stack) MenhirState220 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | LPAREN -> + _menhir_run223 _menhir_env (Obj.magic _menhir_stack) MenhirState220 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | MAPPING -> + _menhir_run221 _menhir_env (Obj.magic _menhir_stack) MenhirState220 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState220) : 'freshtv86) + | LBRACKET -> + _menhir_run5 _menhir_env (Obj.magic _menhir_stack) MenhirState4 _menhir_env._menhir_lexbuf.Lexing.lex_start_p + | EQUAL -> + _menhir_reduce19 _menhir_env (Obj.magic _menhir_stack) MenhirState4 + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + _menhir_errorcase _menhir_env (Obj.magic _menhir_stack) MenhirState4) : 'freshtv88) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv89 * Lexing.position) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv90)) : 'freshtv92) + | OBJECT -> + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv93) = Obj.magic _menhir_stack in + ((let (_, _startpos) = Obj.magic _menhir_stack in + let _endpos = _startpos in + let _v : 'tv_opt_logical_or_trusted = +# 824 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( POnormal ) +# 15333 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_opt_logical_or_trusted _menhir_env _menhir_stack _endpos _v _startpos) : 'freshtv94) + | _ -> + assert (not _menhir_env._menhir_error); + _menhir_env._menhir_error <- true; + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv95 * Lexing.position * 'tv_declarations) = Obj.magic _menhir_stack in + (raise _eRR : 'freshtv96)) : 'freshtv98) + +and _menhir_discard : _menhir_env -> _menhir_env = + fun _menhir_env -> + let lexer = _menhir_env._menhir_lexer in + let lexbuf = _menhir_env._menhir_lexbuf in + let _tok = lexer lexbuf in + { + _menhir_lexer = lexer; + _menhir_lexbuf = lexbuf; + _menhir_token = _tok; + _menhir_error = false; + } + +and file : (Lexing.lexbuf -> token) -> Lexing.lexbuf -> ( +# 207 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + (Parsetree.p_file_structure) +# 15358 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" +) = + fun lexer lexbuf -> + let _menhir_env = + let (lexer : Lexing.lexbuf -> token) = lexer in + let (lexbuf : Lexing.lexbuf) = lexbuf in + ((let _tok = Obj.magic () in + { + _menhir_lexer = lexer; + _menhir_lexbuf = lexbuf; + _menhir_token = _tok; + _menhir_error = false; + }) : _menhir_env) + in + Obj.magic (let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv3) = ((), _menhir_env._menhir_lexbuf.Lexing.lex_curr_p) in + ((let _menhir_env = _menhir_discard _menhir_env in + let (_menhir_env : _menhir_env) = _menhir_env in + let (_menhir_stack : 'freshtv1) = Obj.magic _menhir_stack in + ((let (_, _endpos) = Obj.magic _menhir_stack in + let _v : 'tv_declarations = +# 217 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + ( [] ) +# 15381 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + in + _menhir_goto_declarations _menhir_env _menhir_stack _endpos _v) : 'freshtv2)) : 'freshtv4)) + +# 836 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.mly" + + +# 15388 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" + +# 269 "" + + +# 15393 "/Users/zachpage/Documents/deepsea-2/Edsger/parser.ml" diff --git a/src/Edsger/parser.mli b/src/Edsger/parser.mli new file mode 100644 index 0000000..cf51ec1 --- /dev/null +++ b/src/Edsger/parser.mli @@ -0,0 +1,91 @@ + +(* The type of tokens. *) + +type token = + | XOR + | WITH + | UNEQUAL + | UINT of (string) + | TYPE + | TRUSTED + | TO + | THEN + | STRING of (string) + | STAR + | SLASH + | SIGNATURE + | SHR + | SHL + | SEMICOLON + | RPAREN + | REFINED + | RBRACKET + | RBRACE + | PLUS + | OBJECT + | MOD + | MINUS + | MATCH + | MAPPING + | LPAREN + | LOGICAL + | LIST + | LET + | LESSEQ + | LESS + | LBRACKET + | LBRACE + | LAYER + | INT of (string) + | INDEXED + | IN + | IF + | IDENTITY + | IDENT of (Astcommon.ident) + | GREATEREQ + | GREATER + | GHOST + | FOR + | FOLD + | FIRST + | FAIL + | EXTERNAL + | EVENT + | EQUAL + | EOF + | END + | EMIT + | ELSE + | DOUBLEARROW + | DOT + | DO + | DISJUNCTION + | DENY + | CONSTRUCTOR + | CONST + | CONJUNCTION + | COMMA + | COLONLESS + | COLONGREATER + | COLONCOLON + | COLON + | CLONE + | BITNOT + | BITAND + | BEGIN + | BARBAR + | BAR + | BANG + | AT + | ASSIGN + | ASSERT + | ARROW + | ARRAY + +(* This exception is raised by the monolithic API functions. *) + +exception Error + +(* The monolithic API. *) + +val file: (Lexing.lexbuf -> token) -> Lexing.lexbuf -> (Parsetree.p_file_structure) diff --git a/src/Edsger/parser.mly b/src/Edsger/parser.mly index 58f6c53..1303c5a 100755 --- a/src/Edsger/parser.mly +++ b/src/Edsger/parser.mly @@ -511,21 +511,18 @@ atom: | LPAREN RPAREN { mkexp_ ~loc:$sloc (PEconstant CONunit) } | LPAREN comma_sep_expressions RPAREN { $2 } | a=atom LBRACKET e=expression RBRACKET { mkexp_ ~loc:$sloc (PEindex (a, e)) } - | atom DOT IDENT { mkexp_ ~loc:$sloc (PEfield ($1, $3)) } -; -atoms: - { [] } - | atoms atom { $2 :: $1 } + (* | atom LPAREN atom RPAREN { mkexp_ ~loc:$sloc (PEpair (a1, a2)) } *) + (* | atom DOT IDENT { mkexp_ ~loc:$sloc (PEfield ($1, $3)) } *) ; + expression: - atom atoms - { let x =if $2 = [] - then $1 - else (mkexp_ ~loc:$sloc (PEapp ($1, List.rev $2))) - in - mkexp_ ~loc:$sloc x.p_expression_desc - (* Constructors with zero parameter will be separated from PEglob - in later stages.*) } + | a=atom+ { match a with + | [hd] -> hd + | _ -> mkexp_ ~loc:$sloc (PEapp a) + } + | a1=atom+ DOT a2=atom+ { let x = mkexp_ ~loc:$sloc (PEfield (a1, a2)) in + (* print_endline ("PEfield: " ^ (string_of_p_expression x)); *) x + } | MINUS e=expression %prec prec_unary_prefix { mkexp_ ~loc:$sloc (PEun (OPneg, e)) } | BANG e=expression %prec prec_unary_prefix { mkexp_ ~loc:$sloc (PEun (OPnot, e)) } | BITNOT e=expression %prec prec_unary_prefix { mkexp_ ~loc:$sloc (PEun (OPbitnot, e)) } @@ -797,8 +794,13 @@ layer_slot: ; layer_obj_inst: object_expression { mkobjinst ~loc:$sloc (POinternal $1) } +#ifdef ANT + | IDENTITY LPAREN INT RPAREN COLONLESS object_expression { mkobjinst ~loc:$sloc (POexternal ((CONaddress $3), $6)) } + | IDENTITY LPAREN UINT RPAREN COLONLESS object_expression { mkobjinst ~loc:$sloc (POexternal ((CONaddress $3), $6)) } +#else | ADDRESS LPAREN INT RPAREN COLONLESS object_expression { mkobjinst ~loc:$sloc (POexternal ((CONaddress $3), $6)) } | ADDRESS LPAREN UINT RPAREN COLONLESS object_expression { mkobjinst ~loc:$sloc (POexternal ((CONaddress $3), $6)) } +#endif ; annotations: diff --git a/src/Edsger/parsetree.ml b/src/Edsger/parsetree.ml index c379ce7..217ad06 100755 --- a/src/Edsger/parsetree.ml +++ b/src/Edsger/parsetree.ml @@ -22,9 +22,9 @@ and p_expression_desc = | PEun of unop * p_expression | PEbin of binop * p_expression * p_expression | PEpair of p_expression * p_expression - | PEapp of p_expression * p_expression list + | PEapp of p_expression list | PEstruct of (ident * p_expression) list - | PEfield of p_expression * ident + | PEfield of p_expression list * p_expression list | PEindex of p_expression * p_expression @@ -292,16 +292,19 @@ and string_of_p_expression_desc = function | PEun (op, e) -> string_of_unop op ^ (string_of_p_expression e) | PEbin (op, e1, e2) -> "(" ^ string_of_p_expression e1 ^ " " ^ string_of_binop op ^ " " ^ string_of_p_expression e2 ^ ")" - | PEpair (e1, e2) -> + | PEpair (e1, e2) -> "PEpair:" ^ "(" ^ string_of_p_expression e1 ^ ", " ^ string_of_p_expression e2 ^ ")" - | PEapp (e, es) -> + | PEapp (e::es) -> string_of_p_expression_desc e.p_expression_desc ^ "location info: " ^ string_of_location e.p_expression_loc ^ " (" ^ String.concat ") (" (List.map string_of_p_expression es) ^ ")" | PEstruct ls -> "{" ^ String.concat "; " (List.map (fun (i, e) -> i ^ " = " ^ string_of_p_expression e) ls) ^ "}" - | PEfield (e, i) -> string_of_p_expression e ^ ".{" ^ i ^ "}" + | PEfield (e1, e2) -> "(" ^ String.concat ") (" (List.map string_of_p_expression e1) ^ ")" ^ ".{" ^ + "(" ^ String.concat ") (" (List.map string_of_p_expression e2) + ^ "}" | PEindex (e1, e2) -> string_of_p_expression e1 ^ "[" ^ string_of_p_expression e2 ^ "]" + | PEapp [] -> raise Exit let rec string_of_p_command pcmd = (string_of_p_command_desc pcmd.p_command_desc) (* ^ (string_of_location pcmd.p_command_loc) *) diff --git a/src/Edsger/parsetree.mli b/src/Edsger/parsetree.mli index 425e3a3..d6f8204 100755 --- a/src/Edsger/parsetree.mli +++ b/src/Edsger/parsetree.mli @@ -20,9 +20,9 @@ and p_expression_desc = | PEun of Astcommon.unop * p_expression | PEbin of Astcommon.binop * p_expression * p_expression | PEpair of p_expression * p_expression - | PEapp of p_expression * p_expression list + | PEapp of p_expression list | PEstruct of (Astcommon.ident * p_expression) list - | PEfield of p_expression * Astcommon.ident + | PEfield of p_expression list * p_expression list | PEindex of p_expression * p_expression diff --git a/src/Edsger/typecheck.ml b/src/Edsger/typecheck.ml index 15a61d4..86118b6 100755 --- a/src/Edsger/typecheck.ml +++ b/src/Edsger/typecheck.ml @@ -264,7 +264,7 @@ let parse_constr_annotations ctype_env annos = if e1'.aImplType = ACtint && e2'.aImplType = ACtint then { aImplDesc = ACtimes (e1', e2'); aImplType = ACtint } else raise UnsupportedConstrAnnotation - | PEapp ({p_expression_desc = (PEglob "array_init"); p_expression_loc = _}, [e]) -> (* PEglob "array_init" *) + | PEapp ({p_expression_desc = (PEglob "array_init"); p_expression_loc = _} :: [e]) -> (* PEglob "array_init" *) let e' = parse_expression e in { aImplDesc = ACarray e'; aImplType = ACtarray (0, e'.aImplType) } | PEstruct lst -> @@ -768,6 +768,15 @@ let typecheck parsed filename = ([builtin_type_a_type Tuint],{ aRexprDesc = AEbuiltin ("blockhash", []); aRexprType = builtin_type_a_type Tuint }); ] in + + let get_field_ident e loc = match e with + | hd::_ -> (match hd with + | { p_expression_desc = PEglob s'; p_expression_loc = _ } -> s' + | _ -> report_error ("First atom of field must be ident") + loc; "" + ) + | [] -> report_error ("Unreachable: empty atom list should not be possible") + loc; "" in let get_type i loc = try Hashtbl.find type_environment i @@ -1202,7 +1211,7 @@ let typecheck parsed filename = aBigExprType = t } - | PEapp ({p_expression_desc = (PEglob "keccak256"); p_expression_loc = _ }, [{p_expression_desc = PEpair (e1, e2); p_expression_loc = _ }]) -> + | PEapp ({p_expression_desc = (PEglob "keccak256"); p_expression_loc = _ } :: [{p_expression_desc = PEpair (e1, e2); p_expression_loc = _ }]) -> let e1' = translate_rexpr var_env tmp_env e1 in let e2' = translate_rexpr var_env tmp_env e2 in begin match binop_type OPsha_2 e1'.aRexprType e2'.aRexprType with @@ -1213,7 +1222,7 @@ let typecheck parsed filename = `RExpr { aRexprDesc = AEbinop (OPsha_2, e1', e2'); aRexprType = dummy_type "*BINOP*" } end - | PEapp ({p_expression_desc = (PEglob "keccak256"); p_expression_loc = _}, [e]) -> + | PEapp ({p_expression_desc = (PEglob "keccak256"); p_expression_loc = _} :: [e]) -> let e' = translate_rexpr var_env tmp_env e in begin match unop_type OPsha_1 e'.aRexprType with | Some t -> `RExpr { aRexprDesc = AEunop (OPsha_1, e'); aRexprType = t } @@ -1223,7 +1232,7 @@ let typecheck parsed filename = `RExpr { aRexprDesc = AEunop (OPsha_1, e'); aRexprType = dummy_type "*UNOP*" } end - | PEapp ({p_expression_desc = (PEglob i); p_expression_loc = _}, es) when i = "fst" || i = "snd" -> begin + | PEapp ({p_expression_desc = (PEglob i); p_expression_loc = _} :: es) when i = "fst" || i = "snd" -> begin match es with | [] | _ :: _ :: _ -> report_error (i ^ " only takes one argument; " ^ @@ -1243,7 +1252,7 @@ let typecheck parsed filename = end (* TODO: *) - | PEapp ({p_expression_desc = (PEglob i); p_expression_loc = _}, es) when i = "array_sets" -> begin + | PEapp ({p_expression_desc = (PEglob i); p_expression_loc = _} :: es) when i = "array_sets" -> begin let ct = begin match es with | e::_ -> let e' = translate_big_expr var_env tmp_env e in @@ -1269,7 +1278,7 @@ let typecheck parsed filename = in `BigExpr { aBigExprDesc = AEconstr (c, params); aBigExprType = t } end - | PEapp ({p_expression_desc = (PEglob i); p_expression_loc = _}, es) -> + | PEapp ({p_expression_desc = (PEglob i); p_expression_loc = _} :: es) -> begin match Hashtbl.find_opt ethbuiltins_environment i with | Some (argtypes, {aRexprDesc=AEbuiltin (bi,_); aRexprType=rt}) -> @@ -1397,8 +1406,10 @@ let typecheck parsed filename = `BigExpr dummy_big_expr end - | PEfield (e, f) -> + | PEfield (a1, a2) -> + let e = List.hd a1 in let e' = translate_lexpr var_env tmp_env e in + let f = get_field_ident a2 e.p_expression_loc in begin match e'.aLexprType.aTypeDesc with | ATdata (i, ATsingleton c) -> begin match catch_not_found (List.assoc f) c.aTypeConstrArgs with @@ -1626,11 +1637,34 @@ let typecheck parsed filename = next_id' in match cmd.p_command_desc with (* | PCskip -> ACskip, tvoid_unit, next_id *) - | PCyield ({p_expression_desc = PEapp ({p_expression_desc = (PEfield ({p_expression_desc = (PEglob s); p_expression_loc = _}, m)); p_expression_loc = _ }, arg :: rest); p_expression_loc = _ }) -> - if rest <> [] then - report_warning ("More than one (currying) argument for primitive call " ^ - s ^ "." ^ m ^ ": ignored") cmd.p_command_loc; - begin match find_method env s m with + | PCyield ({ + p_expression_desc = PEfield (a1, a2); + p_expression_loc = _ + }) when (List.length a2 <> 1) (* only process method calls, not struct access *) -> + let s1 = get_field_ident a1 cmd.p_command_loc in (* contract/layer name *) + let s2 = get_field_ident a2 cmd.p_command_loc in (* method name *) + let value, gas = ( (* Both optional. If `value` is given, `gas` must be too. *) + let invalid () = report_error "Invalid value/gas arguments" cmd.p_command_loc; None, None in + match List.length a1 with + | 1 -> None, None + | 2 -> (let e = List.nth a1 1 in + match e.p_expression_desc with + | PEglob v -> Some e, None + | PEpair (v, g) -> ( + match v.p_expression_desc, g.p_expression_desc with + | PEglob _, PEglob _ -> Some v, Some g + | _ -> invalid () + ) + | _ -> invalid () + ) + | _ -> report_error "More than two atoms in field" cmd.p_command_loc; None, None + ) + in + let translate_option = function + | None -> None + | Some v -> Some (translate_rexpr env.variable_env tmp_env v) in + let value', gas' = translate_option value, translate_option gas in + begin match find_method env s1 s2 with | `Not_found msg -> report_error msg cmd.p_command_loc; return next_id tvoid_unit side_effect_pure ACskip | `Found mt -> @@ -1639,8 +1673,11 @@ let typecheck parsed filename = report_error ("Method call in constructor is not working in current version of compiler") cmd.p_command_loc; end; return next_id mt.aMethodReturnType (method_side_effect mt) @@ - ACcall (s, m, unfold_call_arg env.variable_env tmp_env - arg mt.aMethodArgumentTypes) + ACcall (s1, s2, unfold_call_arg env.variable_env tmp_env + (* TODO: catch n too large *) + (* TODO: check that s1 and s2 don't have any extra atoms *) + (List.nth a2 1) mt.aMethodArgumentTypes, + value', gas') end | PCyield ({p_expression_desc = (PEglob v); p_expression_loc = _}) @@ -1659,8 +1696,8 @@ let typecheck parsed filename = return next_id res_t { eff with invokesLogical = true } @@ ACexternal (dest, s, i, []) - | PCyield ({p_expression_desc = PEapp ({p_expression_desc = (PEglob f); p_expression_loc = _ }, arg :: rest); p_expression_loc = _}) - | PCstore (_, {p_expression_desc = PEapp ({p_expression_desc = (PEglob f); p_expression_loc = _ }, arg :: rest); p_expression_loc = _ }) + | PCyield ({p_expression_desc = PEapp ({p_expression_desc = (PEglob f); p_expression_loc = _ } :: (arg :: rest)); p_expression_loc = _}) + | PCstore (_, {p_expression_desc = PEapp ({p_expression_desc = (PEglob f); p_expression_loc = _ } :: (arg :: rest)); p_expression_loc = _ }) when Hashtbl.mem external_function_environment f -> let (s, mt) = Hashtbl.find external_function_environment f in let dest, res_t, eff = match cmd.p_command_desc with @@ -1855,7 +1892,7 @@ let typecheck parsed filename = in return next_id tvoid_unit eff @@ ACconstr (el', er'') end - | PCemit ({p_expression_desc = PEapp ({p_expression_desc = (PEglob f); p_expression_loc = _ }, args :: rest); p_expression_loc = _}) -> begin + | PCemit ({p_expression_desc = PEapp ({p_expression_desc = (PEglob f); p_expression_loc = _ } :: (args :: rest)); p_expression_loc = _}) -> begin if rest <> [] then report_warning ("More than one (currying) argument for event " ^ f ^": ignored") cmd.p_command_loc; diff --git a/src/Edsger/uclidgen.ml b/src/Edsger/uclidgen.ml index 74575a8..e4b3a6b 100644 --- a/src/Edsger/uclidgen.ml +++ b/src/Edsger/uclidgen.ml @@ -232,7 +232,7 @@ let rec str_cmd_uclid acmd base_layer o idnt = idnt ^ str_cmd_uclid c2' base_layer o (idnt ^ " ") ^ "\n" ^ idnt ^ " " ^ x ^ " = ;\n" ^ idnt ^ "}\n" - | ACcall (s, f, es) -> + | ACcall (s, f, es, _, _) -> (match o.aObjectAddress with | Some addr -> let o' = try (List.assoc s base_layer.aLayerAllObjects) @@ -261,7 +261,7 @@ let rec str_cmd_uclid acmd base_layer o idnt = | ACsequence (c1, c2) -> str_cmd_uclid c1 base_layer o idnt ^ str_cmd_uclid c2 base_layer o idnt - | ACcall (s, f, es) -> + | ACcall (s, f, es, _, _) -> (match o.aObjectAddress with | Some addr -> let o' = try (List.assoc s base_layer.aLayerAllObjects) @@ -344,7 +344,7 @@ let rec get_modified_uclid acmd base_layer o modified = SS.union (get_modified_uclid c1 base_layer o modified) (get_modified_uclid c2 base_layer o modified) | ACsequence (c1, c2) -> SS.union (get_modified_uclid c1 base_layer o modified) (get_modified_uclid c2 base_layer o modified) - | ACcall (s, f, es) -> + | ACcall (s, f, es, _, _) -> let o' = try (List.assoc s base_layer.aLayerAllObjects) with Not_found -> raise (PrimitiveNotFound s) in get_modified_uclid (List.find (fun m -> m.aMethodName = f) o'.aObjectMethods).aMethodBody base_layer o' modified @@ -402,7 +402,7 @@ let rec get_vardecl_uclid acmd base_layer o modified = SS.union (get_vardecl_uclid c1 base_layer o mod') (get_vardecl_uclid c2 base_layer o mod') | ACsequence (c1, c2) -> SS.union (get_vardecl_uclid c1 base_layer o modified) (get_vardecl_uclid c2 base_layer o modified) - | ACcall (s, f, es) -> + | ACcall (s, f, es, _, _) -> modified | ACcond (e, c1, c2) -> SS.union (get_vardecl_uclid c1 base_layer o modified) (get_vardecl_uclid c2 base_layer o modified) diff --git a/src/Makefile b/src/Makefile index ed639d1..dc9ee08 100644 --- a/src/Makefile +++ b/src/Makefile @@ -1,16 +1,6 @@ core: core.make make -f core.make -.PHONY: edsger -edsger: Edsger/parser.ml - cd Edsger && dune build edsger.bc -.PHONY: edsger.native -edsger.exe: Edsger/parser.ml - cd Edsger && dune build edsger.exe - -Edsger/parser.ml: Edsger/config.h Edsger/parser.mly Edsger/make_parser.sh - Edsger/make_parser.sh - clean: core.make make -f core.make clean rm -r _build @@ -18,8 +8,14 @@ clean: core.make core.make: _CoqProject coq_makefile -f _CoqProject -o core.make -extraction: - cd backend && coqtop -R .. DeepSpec -top Extract < Extract.v +edsger: Edsger/parser.ml + cd Edsger && dune build edsger.bc + +edsger.exe: Edsger/parser.ml + cd Edsger && dune build edsger.exe + +Edsger/parser.ml: Edsger/config.h Edsger/parser.mly Edsger/make_parser.sh + cd Edsger && ./make_parser.sh minicc: cd minic && dune build minicc.bc @@ -27,4 +23,4 @@ minicc: minicc.exe: cd minic && dune build minicc.exe -.PHONY: clean core extraction minicc minicc.exe +.PHONY: clean core edsger edsger.exe diff --git a/src/Runtime.v b/src/Runtime.v new file mode 100755 index 0000000..e1ec9b2 --- /dev/null +++ b/src/Runtime.v @@ -0,0 +1,860 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** The run-time library for certified programs generated by [edsger] *) + +(* Standard library modules *) +Require Export BinInt. + +(* CompCert modules *) +Require Export backend.Cop. +Require Export backend.Ctypes. +Require Export cclib.Coqlib. +Require Export cclib.Integers. +Require Export cclib.Maps. + +(* CertiKOS modules *) +(* +Require liblayers.compat.CompatGenSem. +Require liblayers.compcertx.MakeProgram. +Require liblayers.compcertx.StencilImpl. +Require mcertikos.layerlib.PrimSemantics. + *) + +(* DeepSpec modules *) +Require Export DeepSpec.core.Cval. +Require Export DeepSpec.core.SEnv. +Require Export DeepSpec.core.HyperType. +Require Export DeepSpec.core.HyperTypeInst. +Require Export DeepSpec.core.MemoryModel. +Require Export DeepSpec.core.Syntax. +Require Export DeepSpec.core.SynthesisFunc. +Require Export DeepSpec.core.SynthesisStmt. +Require Export DeepSpec.lib.OProp. +(* Todo: do we need this export? *) +Require Export DeepSpec.lib.IndexedMaps. +Require Export DeepSpec.lib.SimpleMaps. +Require Export DeepSpec.lib.SimpleIndexedMaps. + +(* We don't import liblayers.logic.Structures, which populates the notation + scope too much; instead we declare the minimum we need. *) +(* +Infix "⊕" := liblayers.logic.Structures.oplus (at level 60, right associativity). +Infix "↦" := liblayers.logic.Structures.mapsto (at level 55, right associativity). +*) +Notation tchar := (Tint I8 Unsigned). + +(* This file will contain definitions related to each type in DeepSpec, + defining the corresponding lenses. *) + +(* Default value for global pointers. + + The current version of DeepSpec does not provide pointers as a primitive type, but + if we added an address-of operator, then this would be used. *) +(* +Definition empty_globalpointer := GLOBUndef. + *) + +Lemma and_iff_compat {A B C D} : (A <-> C) -> (B <-> D) -> ((A /\ B) <-> (C /\ D)). +Proof. + intros A_iff_C B_iff_D. + rewrite A_iff_C, B_iff_D. + apply iff_refl. +Qed. + +Ltac solve_algebraic_record_data_type_hyper_type hyper_type_impl := + esplit; + [ intros; eexists; reflexivity + | unfold ht_ft_cond, ht_default, hyper_type_impl; + repeat match goal with + | |- _ /\ _ => split + | |- True => exact I + end; apply ht_default_ft_cond + | intros f; apply iff_refl || ( + unfold ht_valid_ft_cond, ht_valid_ft_ocond, hyper_type_impl; + rewrite ? oand1_distr, ? OProp1map1; [| exact I .. ]; + repeat (apply and_iff_compat; try apply @ht_valid_ft_ocond_same); + typeclasses eauto) + ]. + +Ltac solve_record_type_hyper_field := + constructor; + [ reflexivity (* size_pos *) + | simpl; tauto (* get_returns *) + | simpl; tauto (* set_returns *) + | simpl; intros; constructor; reflexivity (* get_correct*) + | intros part whole pcond wcond; constructor; reflexivity (* set_correct *) + | eexists; eexists; + repeat apply conj; reflexivity (* delta_correct *) + ]. + +(* Previous version. Still unclear on this... *) +(* Ltac solve_record_type_hyper_field := + esplit; + [ reflexivity + | simpl; auto + | intros whole wcond; apply wcond + | intros part whole pcond wcond; simpl; repeat apply conj; solve [ apply pcond | apply wcond ] + | constructor; reflexivity + | intros part whole pc wc; apply ht_some_cval; reflexivity + | eexists; eexists; + repeat apply conj; reflexivity + ]. *) + +Hint Extern 1 ((_ < _)%Z) => reflexivity : zarith. +Hint Extern 1 ((_ > _)%Z) => reflexivity : zarith. + +(* TODO: Oops, the currently admitted condition is supposed to be available from the context, we need to track down where it's supposed to come from. *) +Ltac solve_twobranch_type_hyper_field := + esplit; + [ simpl; constructor (* size_pos *) + | intros whole wcond; + try destruct whole as [ | ]; simpl in *; try tauto; auto with zarith; apply wcond (* get_returns *) + | intros part whole pcond wcond; simpl; try destruct whole as [ | ]; repeat apply conj; solve [ apply pcond | apply wcond ] (* set_returns *) + | intros whole; + try destruct whole as [ | ]; simpl in *; try tauto; + constructor; reflexivity (* get_correct *) + | simpl; intros part whole; + assert (wvalid : ht_valid_ft_cond whole) by admit; + try destruct whole as [ | ]; simpl in wvalid; try tauto; + constructor; reflexivity (* set_correct *) + | eexists; eexists; + repeat apply conj; reflexivity (* delta_correct *) + ]. + +Definition array_cval {tp} `{HyperTypeImpl tp} arr (size : Z) := + CVarray (*unpair_ty tp*) (CAmap (ZMap.map (@ht_cval tp _) arr)). + +Lemma forall_suchthat_iff_compat {T P A B} : + (forall t : T, P t -> (A t <-> B t)) -> + ((forall t : T, P t -> A t) <-> (forall t : T, P t -> B t)). +Proof. + intros IH; split; intros other t p; + [ rewrite <- IH | rewrite IH ]; try apply other; exact p. +Qed. + +Lemma forall_suchthat_iff_compat2 {T S P A B} : + (forall (t : T) (s:S), P t s -> (A t s <-> B t s)) -> + ((forall (t : T) (s:S), P t s -> A t s) <-> (forall (t : T) (s:S), P t s -> B t s)). +Proof. + intros IH; split; intros other t s p; + [ rewrite <- IH | rewrite IH ]; try apply other; exact p. +Qed. + + +Open Scope Z_scope. + +Lemma zle_and_zlt_between a b c : a <= b < c -> zle a b && zlt b c = true. +Proof. + intros [ a_le_b b_lt_c ]. + destruct (zle a b); try contradiction. + destruct (zlt b c); try contradiction. + reflexivity. +Qed. + +(* Lenses to do with arrays. *) +Section ZMAP_BASED_ARRAY. + Variable tp : type_pair. + (* Hypothesis tp_naturally_aligned : naturally_aligned (unpair_ty tp). *) + Context`{hti : HyperTypeImpl tp}. + Variable sz : Z. + + Let arr_tp := Tpair (ZMap.t (unpair_ft tp)) (Tarray (unpair_ty tp) sz). + + Local Instance ZMap_based_array_hyper_type_impl : HyperTypeImpl arr_tp := { + ht_cval f := array_cval (*tp*) f sz; + ht_ft_cond f := forall n, 0 <= n < sz -> @ht_ft_cond tp _ (ZMap.get n f); + ht_default := ZMap.init (@ht_default tp _); + ht_valid_ft_cond f := forall n, 0 <= n < sz -> + @ht_valid_ft_cond tp _ (ZMap.get n f); + ht_valid_ft_ocond := + omap1 (fun p f => forall n, 0 <= n < sz -> p (ZMap.get n f)) + (@ht_valid_ft_ocond tp _); +(* ht_inject j a b := forall n, (*0 <= n < sz ->*) (* strengthened for passthrough *) + @ht_inject tp _ j (ZMap.get n a) (ZMap.get n b) *) + }. + + Local Instance ZMap_based_array_hyper_type`{ht : !HyperType tp} : HyperType arr_tp. + Proof. esplit. + - intros; eexists; reflexivity. + - intros n n_range. + unfold ht_default, ZMap_based_array_hyper_type_impl. + rewrite ZMap.gi. + apply ht_default_ft_cond. + - intros f. + unfold ht_valid_ft_cond, ht_valid_ft_ocond, ZMap_based_array_hyper_type_impl. + change (ZMap.t (unpair_ft tp)) with (unpair_ft arr_tp). + rewrite OProp1map1; [| intros; exact I ]. + apply forall_suchthat_iff_compat. + intros; apply ht_valid_ft_ocond_same. + Qed. + + Local Instance ZMap_based_array_hyper_index_impl : HyperIndexImpl arr_tp tp := { + Hindex_size := sz; + Hindex_get idx ar := ZMap.get idx ar; + Hindex_set idx elem ar := ZMap.set idx elem ar + }. + + Local Instance ZMap_based_array_hyper_index`{ht : !HyperType tp} : HyperIndex arr_tp tp. + Proof. Admitted. +(* Proof. esplit. *) +(* - (* Hindex_get_returns *) *) +(* intros a i i_range ac; apply ac, i_range. *) +(* - (* Hindex_set_returns *) *) +(* simpl; intros f a i i_range fc ac i' i'_range. *) +(* destruct (zeq i' i) as [ -> | i_ne ]; *) +(* [ rewrite ? ZMap.gss; apply fc *) +(* | rewrite ? ZMap.gso; [ apply ac, i'_range | apply i_ne ] *) +(* ]. *) +(* - (* Hindex_get_correct *) *) +(* simpl; intros a i i_range. *) +(* rewrite zle_and_zlt_between, ZMap.gmap; try assumption. *) +(* constructor. *) +(* reflexivity. *) +(* - (* Hindex_set_correct *) *) +(* simpl; intros f a i i_range. *) +(* rewrite zle_and_zlt_between; try assumption. *) +(* constructor. *) +(* simpl; unfold array_cval. *) +(* do 2 f_equal. *) +(* apply eq_sym, ZMapAux.smap. *) +(* - (* Hindex_array_type *) *) +(* eexists; reflexivity. *) +(* (* - (* Hindex_naturally_aligned *) *) +(* exact tp_naturally_aligned. *) *) +(* Qed. *) + + Local Instance ZMap_based_array_hyper_index_passthrough : + HyperIndexPassthrough arr_tp tp. +(* Proof. esplit; simpl. + - intros j a b. + exact (fun x => x). + - intros j elem_a arr_a elem_b arr_b elem_inj arr_inj i n. + destruct (zeq n i) as [ -> | i_ne ]. + + rewrite 2 ZMap.gss. + exact elem_inj. + + rewrite 2 ZMap.gso; try assumption. + apply arr_inj. + Qed. *) +End ZMAP_BASED_ARRAY. + + +Definition hashmap_cval {tp} `{HyperTypeImpl tp} arr := + CVhashmap + (CHmap (map_of_tree ht_cval ht_default arr)). + +(* Lenses to do with hashmaps. *) +Section INT256TREE_BASED_INT_HASHMAP. + Variable tp : type_pair. + (* Hypothesis tp_naturally_aligned : naturally_aligned (unpair_ty tp). *) + Context`{hti : HyperTypeImpl tp}. + + Let arr_tp := Tpair (Int256Tree.t (unpair_ft tp)) (Thashmap tint (unpair_ty tp)). + + Local Instance Int256Tree_based_hashmap_hyper_type_impl : HyperTypeImpl arr_tp := { + ht_cval f := hashmap_cval (*tp*) f; + ht_ft_cond f := forall n v, Int256Tree.get n f = Some v -> @ht_ft_cond tp _ v; + ht_default := (Int256Tree.empty (unpair_ft tp)); + ht_valid_ft_cond f := forall n v, Int256Tree.get n f = Some v -> + @ht_valid_ft_cond tp _ v; + ht_valid_ft_ocond := + omap1 (fun p f => forall n v, Int256Tree.get n f = Some v -> p v) + (@ht_valid_ft_ocond tp _); + }. + + Local Instance Int256Tree_based_hashmap_hyper_type`{ht : !HyperType tp} : HyperType arr_tp. + Proof. esplit. + - intros; eexists; reflexivity. + - intros n. + unfold ht_default, Int256Tree_based_hashmap_hyper_type_impl. + rewrite Int256Tree.gempty. + intros; congruence. + - intros f. + unfold ht_valid_ft_cond, ht_valid_ft_ocond, Int256Tree_based_hashmap_hyper_type_impl. + change (Int256Tree.t (unpair_ft tp)) with (unpair_ft arr_tp). + rewrite OProp1map1; [| intros; exact I ]. + apply forall_suchthat_iff_compat2. + intros; apply ht_valid_ft_ocond_same. + Qed. + + Local Instance Int256Tree_based_hashmap_hash_impl : HyperIntHashImpl arr_tp tp := { + Hhash_get idx ar := Int256Tree.get_default ht_default idx ar; + Hhash_set idx elem ar := Int256Tree.set idx elem ar + }. + + Require Import lib.ArithInv. + + Local Instance Int256Tree_based_hashmap_hash `{ht : !HyperType tp} : HyperIntHash arr_tp tp. + Proof. esplit. + - (* Hhash_get_returns *) + intros a i ac. + simpl. + unfold Int256Tree.get_default. + destruct (Int256Tree.get i a) eqn:?. + + apply (ac _ _ Heqo). + + apply ht_default_ft_cond. + - (* Hhash_set_returns *) + simpl; intros f a i fc ac i'. + intros. + destruct (Int256.eq i' i) eqn:i_eq; inv_arith. + + subst. + rewrite Int256Tree.gss in H. + inversion H. + subst; auto. + + rewrite Int256Tree.gso in H by assumption. + eapply ac; eauto. + - (* Hhash_get_correct *) + simpl; intros a i. + unfold Int256Tree.get_default. + destruct (Int256Tree.get i a) as [v|] eqn:?. + + eapply (map_of_tree_Some ht_cval ht_default) in Heqo. + rewrite Heqo. + constructor. + reflexivity. + + apply (map_of_tree_None ht_cval ht_default) in Heqo. + rewrite Heqo. + constructor. + reflexivity. + - (* Hhash_set_correct *) + simpl; intros f a i. + constructor. + simpl. + unfold hashmap_cval. + rewrite map_of_tree_set. + reflexivity. + - (* Hhash_array_type *) + eexists; reflexivity. + Qed. + +End INT256TREE_BASED_INT_HASHMAP. + + +(* Recall that layer interfaces are specified in terms of a Coq relations (the "semantics"), + while the desugaring uses Coq functions. + + These functions convert desugared values into the semantic values that the layer interface specifications expect. + *) + +(* Omitting all of this for now... *) +(* +Section SEMOF_HYPER_TYPE. + Context {mem}`{Hmem: Mem.MemoryModel mem}`{Hmwd: MemWithData.UseMemWithData mem}. + Context {data : Type}. + Context`{LayerSpec : !LayerSpecClass}. + Context {ft ty}`{ht : HyperType (Tpair ft ty), hta : !HyperArgRet (Tpair ft ty)}. + + Let unfold_cval cv := match cv with + | CVval v => v + | _ => Vundef + end. + + (* Roughly, this typeclass instance says how to give a relational semantics for desugared functions of type + (data -> option (data * ft)) + in terms of a layerlib relation (which involves zero arguments, a return type ty, and implicitly some state d). + *) + Inductive semof_nil_ft {data} + : GenSem.Semof data (data -> option (data * ft)) Tnil ty := + | semof_nil_ft_intro f (d : data) d' a v: + f d = Monad.ret (d', a) -> + ht_ft_cond a -> ht_valid_ft_cond a -> + unfold_cval (ht_cval a) = v -> + GenSem.semof f nil d v d'. + + Global Existing Instance semof_nil_ft. + + (* Each semantic needs to satisfy some extra properties, so we prove those. *) + Global Instance semof_nil_ft_props + : GenSem.Semprops (data -> option (data * ft)). + Proof. + split. + + (* semprops_well_typed *) + intros ? ? ? ? ? H. + inv H. + assert (ht_has_type := ht_has_type). + eapply ht_has_type; try eassumption. + unfold ht_rel. + exploit ht_basic; try eassumption; inversion 1; + constructor; reflexivity. + + (* semprops_arity *) + intros ? ? ? ? ? H. + inv H; + reflexivity. + + (* semprops_lessdef *) + intros ? ? ? ? ? ? H Hl. + inv H; + inv Hl; + reflexivity. + + (* semprops_inject_neutral *) + inversion 1; subst. + exploit ht_basic; try eassumption; inversion 1; constructor. + + (* semprops_determ *) + inversion 1. + inversion 1. + unfold Monad.bind, Monad.ret in *; simpl in *. + split; [ assert (a0 = a) |]; try congruence. + + (* semprops_inject *) + inversion 1. + inversion 1. + reflexivity. + Qed. + + (* We use the above definition to give semantics to pure functions (which do not change the abstract data). *) + Local Notation lift_nil_ft_pure f := + (fun d => Monad.bind (f d) (fun b => Monad.ret (d, b))). + + Global Instance semof_nil_ft_pure + : GenSem.Semof data (data -> option ft) Tnil ty := + fun f => GenSem.semof (lift_nil_ft_pure f). + + Global Instance semof_nil_ft_pure_props + : GenSem.Semprops (data -> option ft). + Proof. + split; intro f. + + exact (GenSem.semprops_well_typed (lift_nil_ft_pure f)). + + exact (GenSem.semprops_arity (lift_nil_ft_pure f)). + + exact (GenSem.semprops_lessdef (lift_nil_ft_pure f)). + + exact (GenSem.semprops_inject_neutral (lift_nil_ft_pure f)). + + exact (GenSem.semprops_determ (lift_nil_ft_pure f)). + + exact (GenSem.semprops_inject (lift_nil_ft_pure f)). + Qed. + + Global Instance semof_nil_ft_pure_invar `{CompatData data}(f: data -> option ft): + CompatGenSem.PreservesInvariants f. + Proof. + split. + * intros ? ? ? ? ? semof inv. + simpl in semof. + inversion semof as [ ? ? ? ? ? spec ]; subst. + unfold Monad.bind, Monad.ret in spec; simpl in spec. + destruct (f d); try discriminate spec. + injection spec as _ <-; assumption. + * intros ? ? ? ? semof inv. + simpl in semof. + inversion semof as [ ? ? ? ? ? spec ]; subst. + unfold Monad.bind, Monad.ret in spec; simpl in spec. + destruct (f d); try discriminate spec. + injection spec as _ <-; assumption. + * intros ? ? ? ? semof inv. + simpl in semof. + inversion semof as [ ? ? ? ? ? spec ]; subst. + unfold Monad.bind, Monad.ret in spec; simpl in spec. + destruct (f d); try discriminate spec. + injection spec as _ <-; assumption. + Qed. + + (* Then we give semantics for functions that take more than zero arguments. + We use some "generic programming" to specify this in terms of prepending one argument at a time. + *) + Inductive semof_cons_ft `{GenSem.Semof data} + : GenSem.Semof data (ft -> T) (Tcons ty targs) tres := + | semof_cons_ft_intro f a l d v d' va: + GenSem.semof (f a) l d v d' -> + ht_ft_cond a -> ht_valid_ft_cond a -> + unfold_cval (ht_cval a) = va -> + GenSem.semof f (va :: l) d v d'. + + Global Existing Instance semof_cons_ft. + + Global Instance semof_cons_ft_props {T}`(HT: GenSem.Semprops data T) + : GenSem.Semprops (ft -> T). + Proof. + split. + + (* semprops_well_typed *) + intros ? ? ? ? ? H. + inv H. + apply (GenSem.semprops_well_typed (targs := targs) (f a) l d vres d'). + eassumption. + + (* semprops_arity *) + intros ? ? ? ? ? H. + inv H. + simpl. + f_equal. + eapply GenSem.semprops_arity. + eassumption. + + (* semprops_lessdef *) + intros until d'. + intros H Hl. + inv H. + inv Hl. + inv H4. + * f_equal. + eapply GenSem.semprops_lessdef; eassumption. + * revert H3. + exploit ht_basic; try eassumption; inversion 1; discriminate. + + (* semprops_inject_neutral *) + inversion 1; subst. + eapply GenSem.semprops_inject_neutral. + eassumption. + + (* semprops_determ *) + inversion 1. + inversion 1. + assert (a0 = a). + { apply ht_injective; try assumption. + revert H3 H19. + exploit (ht_basic a0); try eassumption. + exploit (ht_basic a); try eassumption. + inversion 1; inversion 1; simpl; congruence. + } + subst. + eapply GenSem.semprops_determ; eassumption. + + (* semprops_inject *) + intros. + inv H. + inv H0. + (*inv H4; try (destruct b; discriminate).*) + f_equal. + - revert H5. + exploit ht_basic; try eassumption; inversion 1; inversion 1; reflexivity. + - eapply GenSem.semprops_inject; eassumption. + Qed. + + Global Instance semof_cons_ft_invar + `{Tsemof: GenSem.Semof data, CompatData data} (f: ft -> T): + (forall n, CompatGenSem.PreservesInvariants (f n)) -> + CompatGenSem.PreservesInvariants f. + Proof. + intros Hf. + split. + * intros ? ? ? ? ? semof inv. + simpl in semof. + inversion semof as [ ? ? ? ? ? ? ? semof' ]; subst. + eapply CompatGenSem.semprops_low_level_invariant; eassumption. + * intros ? ? ? ? semof inv. + simpl in semof. + inversion semof as [ ? ? ? ? ? ? ? semof' ]; subst. + eapply CompatGenSem.semprops_high_level_invariant; eassumption. + * intros ? ? ? ? semof inv. + simpl in semof. + inversion semof as [ ? ? ? ? ? ? ? semof' ]; subst. + eapply CompatGenSem.semprops_kernel_mode; eassumption. + Qed. +End SEMOF_HYPER_TYPE. + +Remove Hints GenSem.semof_nil_int64_pure_total + GenSem.semof_nil_int64_pure + GenSem.semof_nil_int64 + GenSem.semof_nil_int_pure_total + GenSem.semof_nil_int_pure + GenSem.semof_nil_int + GenSem.semof_cons64 + GenSem.semof_cons : typeclass_instances. + +(* These two tactics convert between layerlib semantics and the DeepSpec desugared + representation. *) +Ltac inv_semof' sem_of := + repeat + match type of sem_of with + | GenSem.semof _ _ _ _ _ => unfold GenSem.semof in sem_of + | semof_nil_ft_pure _ _ _ _ _ => unfold semof_nil_ft_pure in sem_of + | semof_cons_ft _ _ _ _ _ => + let a := fresh "arg" in + let sem_of' := fresh "sem_of'" in + let ac := fresh "argc" in + let a_valid := fresh "arg_valid" in + let a_eq := fresh "arg_eq" in + inversion sem_of as [ ? a ? ? ? ? ? sem_of' ac a_valid a_eq ]; subst; + clear sem_of; rename sem_of' into sem_of + | semof_nil_ft _ _ _ _ _ => + let r := fresh "ret" in + let sem_of' := fresh "sem_of'" in + let rc := fresh "retc" in + let r_valid := fresh "ret_valid" in + let r_eq := fresh "ret_eq" in + inversion sem_of as [ ? ? ? r ? sem_of' rc r_valid r_eq ]; subst; + clear sem_of; rename sem_of' into sem_of + | GenSem.semof_nil_void _ _ _ _ _ => + let sem_of' := fresh "sem_of'" in + inversion sem_of as [ ? ? ? sem_of' ]; subst; + clear sem_of; rename sem_of' into sem_of + | Monad.bind ?f ?m = Monad.ret ?b => + let sem_of' := fresh "sem_of'" in + let H := fresh "H" in + destruct (Monad.monad_inv_bind f m b sem_of) as (? & sem_of' & H); + injection H as -> <-; clear sem_of H; rename sem_of' into sem_of + end. + +Ltac inv_semof sem_of := + inv_semof' sem_of; + simpl in sem_of; + try (match type of sem_of with + match ?X with _ => _ end = Some _ => + let Hsem_of := fresh sem_of in + destruct X eqn:Hsem_of; + [ let H'sem_of := fresh Hsem_of in + rename Hsem_of into H'sem_of; rename sem_of into Hsem_of; rename H'sem_of into sem_of + | discriminate] + end). + +Ltac construct_semof := + repeat match goal with + | |- GenSem.semof _ _ _ _ _ => unfold GenSem.semof + | given : ?f _ = Monad.ret ?r |- semof_nil_ft_pure ?spec _ _ _ _ => + match spec with + | context[f] => apply semof_nil_ft_intro with r; try assumption; try reflexivity + end + | given : appcontext[?f ?arg] |- semof_cons_ft ?f _ _ _ _ => + apply semof_cons_ft_intro with arg; try assumption; try reflexivity + | given : ?f ?d = Some (_, ?r) |- semof_nil_ft ?f _ ?d _ _ => + apply semof_nil_ft_intro with r; try assumption; try reflexivity + | |- GenSem.semof_nil_void _ _ _ _ _ => constructor + | |- _ = Monad.ret _ => eassumption + end. + +(* Then we construct the liblayer semantics for primitives. + (This is the right hand side of each entry in a layer interface.) + + There already is some code in liblayer for building primitive semantics, + but it does it in one step directly from Coq functions. Here we do it in + two steps, + coq function -> c semantics -> arbitrary primitive semantics. + This is the first step, the second step is basically just inl(-). +*) +Section primitive_primsem. + Context {mem}`{memory_model : Mem.MemoryModel mem}. + Context`{Hmwd : !MemWithData.UseMemWithData mem}. + Context {D}`{HD : CompatData D}. + + Definition gencsem `{Tsemof : GenSem.Semof (cdata D)}`{Hsemof : !GenSem.Semprops T} + (f : T){H : CompatGenSem.PreservesInvariants f} := {| + CompatCPrimitives.sextcall_primsem_step := CompatGenSem.sextcall_generic_info f; + CompatCPrimitives.sextcall_props := Errors.OK (CompatGenSem.extcall_generic_properties f); + CompatCPrimitives.sextcall_invs := Errors.OK (CompatGenSem.extcall_generic_invariants f H) + |}. + + Definition dummy_sextcall_primsem : CompatCPrimitives.sextcall_primsem (cdata D) := {| + CompatCPrimitives.sextcall_primsem_step := {| + CompatCPrimitives.sextcall_step WB vargs m vres m' := True; + CompatCPrimitives.sextcall_csig := {| + CompcertStructures.csig_args := Ctypes.Tnil; + CompcertStructures.csig_res := Ctypes.Tvoid; + CompcertStructures.csig_cc := AST.cc_default + |} + |}; + CompatCPrimitives.sextcall_props := Errors.Error nil; + CompatCPrimitives.sextcall_invs := Errors.Error nil + |}. +End primitive_primsem. +*) + +(* Each DeepSpec project will generate one instance of this class. + + Then the Coq typeclass search will produce an instance of LayerSpecClass, + which is used by most of synthesis functions. + *) + +Definition block := positive. (* TODO: make a better memory model. *) +Class GlobalAbData {AbData : Type}(AbData_empty_data : AbData) + (AbData_low_level_invariant : block -> AbData -> Prop) := { + AbData_low_level_invariant_incr : forall n n' d, + (n <= n')%positive -> + AbData_low_level_invariant n d -> + AbData_low_level_invariant n' d; + AbData_empty_data_low_level_invariant : forall n, + AbData_low_level_invariant n AbData_empty_data +}. + +(* This is an empty layer with no objects and no import/exports. + The lowest layer defined by the user will build on top of this. *) +Section BUILTIN_LAYER. + Context`{LayerSpec : LayerSpecClass}. + Context {AbData_empty_data : GetHighData}. + Context {AbData_low_level_invariant : block -> GetHighData -> Prop}. + Context`{global_abdata : !GlobalAbData AbData_empty_data AbData_low_level_invariant}. + + Instance BuiltinBase_data_ops : CompatDataOps GetHighData := { + empty_data := AbData_empty_data; + high_level_invariant d := True; + (*low_level_invariant := AbData_low_level_invariant; + kernel_mode d := True *) + }. + + Instance BuiltinBase_data : CompatData GetHighData := { +(* low_level_invariant_incr := AbData_low_level_invariant_incr; + empty_data_low_level_invariant := AbData_empty_data_low_level_invariant; *) + empty_data_high_level_invariant := I + }. + + Definition BuiltinBase_cdata : compatdata := cdata GetHighData. + + (* + Definition BuiltinBase_Layer : liblayers.compat.CompatLayerDef.compatlayer BuiltinBase_cdata + := Structures.emptyset. *) +End BUILTIN_LAYER. + +(* We start numbering identifiers from 10 instead of 1. + + This is because CompCert uses some low-numbered idenfiers for special variables, and + when printing error messages it will print hardcoded names for these instead of numbers(!), + which would be very confusing. +*) +Notation BuiltinBase_local_ident_start := 10%positive. + +(* Tactics to solve the _wellformed and _prf conditions. + Edsger will call these in the generated LayerFoo.v files. +*) +Ltac solve_wellformed := + compute -[Int256.repr Int256.modulus not]; + repeat match goal with + | |- prod ?X ?Y => split + | |- SimpleIndexedMaps.Isomorphism ?X ?Y => + (* simpl may still leave some symbols folded, definitional + equivalent terms may not be syntactically identical. *) + apply SimpleIndexedMaps.identity_automorphism + | |- True => exact I + | |- ?X <> ?Y => discriminate + end; + match goal with + | |- ?g => idtac "unknown well-formed condition" g + end. + +(* +Ltac solve_passthrough_prf f := + unfold f; + repeat match goal with + | |- variable_passthrough_prf _ => assumption + | |- expr_constr_passthrough_prf_conj _ => + unfold expr_constr_passthrough_prf_conj; simpl + | |- { _ : _ | _ } => econstructor + | _ => typeclasses eauto + | |- cmd_constr_passthrough_prf _ (CCrespec_opt _ _ _) => + constructor; intros; CommonTactic.subdestruct; eexists; intros; reflexivity + (* + | |- function_constr_passthrough_prf _ => constructor + | |- cmd_constr_passthrough_prf _ _ => constructor + | |- lexpr_constr_passthrough_prf _ => constructor + | |- expr_constr_passthrough_prf _ => constructor + *) + | _ => constructor + end. + *) + +Ltac solve_function_constr_prf f := + unfold f; + esplit; [ simpl; typeclasses eauto (* HyperArgRet *) || exact I |]; + repeat match goal with + | |- variable_prf _ => constructor; [ reflexivity | typeclasses eauto ] + | |- expr_constr_prf_conj _ => unfold expr_constr_prf_conj; simpl + | |- @ht_ft_cond _ int_Z32_impl _ => split; reflexivity + | |- @ht_ft_cond _ int_bool_impl _ => simpl + (* + | |- primitive_prf _ => typeclasses eauto + | |- primitive_passthrough_prf _ => typeclasses eauto + | |- HyperField _ _ _ => typeclasses eauto + | |- HyperIndex _ _ => typeclasses eauto + | |- HyperByValueType _ _ => typeclasses eauto + | |- HyperBinaryOp _ _ _ _ => typeclasses eauto + *) + | |- _ => typeclasses eauto + (* + | |- cmd_constr_prf _ _ _ _ => econstructor + | |- lexpr_constr_prf _ => econstructor + | |- expr_constr_prf _ => econstructor + | |- True => constructor + | |- _ = _ => constructor + | |- _ /\ _ => constructor + | |- { _ : _ | _ } => econstructor + *) + | |- _ => econstructor + end. + +(* "safety condition destruct". This unfolds synth_stmt_spec_cond. *) +Ltac scdestruct sc := + match type of sc with + | True => clear sc + | ?X = ?X => clear sc + | _ /\ _ => + let sc1 := fresh "sc1" in + let sc2 := fresh "sc2" in + destruct sc as [ sc1 sc2 ]; + scdestruct sc1; scdestruct sc2 + | match ?X with _ => _ end = Some _ => + let sc1 := fresh "sc" in + destruct X eqn:sc1; + try discriminate sc; + scdestruct sc1 (* ; scdestruct sc *) + | _ => progress hnf in sc; scdestruct sc + | _ => idtac + end. + +(* These unfold just enough definitions so that the verification conditions can be discharged. *) +Ltac vcgen_static eval_in := + match goal with + | |- _ /\ _ => split; vcgen_static eval_in + | |- True => exact I + | |- forall X, @?P X -> _ => + let X' := fresh X in + let P' := fresh X' "_cond" in + intros X' P'; + eval_in P'; + vcgen_static eval_in + | |- _ -> False => idtac + | |- _ -> _ => + let sc := fresh "sc" in + intros sc; + eval_in sc; + scdestruct sc; vcgen_static eval_in + | |- let X := _ in _ => + let X' := fresh X in + intros X'; + eval_in X'; + vcgen_static eval_in + | _ => idtac + end. +(* Ltac vcgen' eval_in eval_goal steps := *) +(* vcgen_static eval_in; *) +(* match goal with *) +(* | |- Z.lt _ _ => eval_goal *) +(* | |- Z.gt _ _ => eval_goal *) +(* | |- Z.le _ _ => eval_goal *) +(* | |- Z.ge _ _ => eval_goal *) +(* | |- not _ => eval_goal *) +(* | |- synth_stmt_cond ?c _ _ _ _ _ => *) +(* match c with *) +(* | CCskip => idtac *) +(* | CCload _ => idtac *) +(* | CCyield _ => idtac *) +(* | CCpanic => idtac *) +(* end; *) +(* unfold synth_stmt_cond; eval_goal *) +(* | _ => match steps with *) +(* | S ?steps' => progress hnf; vcgen' eval_in eval_goal steps' *) +(* | _ => idtac *) +(* end *) +(* end. *) + +(* Ltac cbv_minus := cbv -[Int256.repr Int256.modulus zeq zle zlt Z.iter Z.le Z.lt Z.gt Z.ge Z.eqb Z.leb Z.ltb Z.geb Z.gtb Z.mul Z.div Z.modulo Z.add Z.sub Z.shiftl Z.shiftr Z.lxor Z.land Z.lor *) +(* is_true bool_dec ZMap.get ZMap.set hlist_hd *) +(* omap2 oand2 oimply2 oabsorption2 hlist_param_func SpecTree.get SpecTree.get_eq]. *) + +Lemma repr_Z_mod_modulus_eq: forall x, Int.repr (Int.Z_mod_modulus x) = Int.repr x. +Proof. + intro x. + Transparent Int.repr. + unfold Int.repr. + Local Opaque Int.repr. + apply Int.mkint_eq. + rewrite Int.Z_mod_modulus_eq. + apply Zmod_small. + apply Int.Z_mod_modulus_range. +Qed. diff --git a/src/_CoqProject b/src/_CoqProject new file mode 100644 index 0000000..bf1225c --- /dev/null +++ b/src/_CoqProject @@ -0,0 +1,74 @@ +-R . DeepSpec + +-arg -w +-arg none + +cclib/Integers.v +cclib/Coqlib.v +backend/AST.v + +cclib/Maps.v +cclib/Errors.v +lib/Monad/Functor.v +lib/Monad/Applicative.v +lib/Monad/MonadZero.v +lib/Monad/MonadPlus.v +lib/Monad/MonadTrans.v +lib/Monad/MonadReader.v +lib/Monad/MonadExc.v +lib/Monad/OptErrMonadLaws.v +lib/Monad/OptErrMonad.v +lib/Monad/OptionMonadLaws.v +lib/Monad/MonadState.v +lib/Monad/MonadLaws.v +lib/Monad/MonadInv.v +lib/Monad/OptionMonad.v +lib/Monad/RunStateTInv.v +backend/Smallstep.v +backend/Events.v +backend/phase/MiniC/BigstepSemantics.v +backend/Environments/HashEnv.v +backend/Environments/StackEnv.v +backend/MachineModelLow.v +backend/Expressions/ExpStacked.v +backend/GasModel.v +backend/phase/MiniC/Semantics.v +backend/TempModel.v +backend/Expressions/SemanticsMiniC.v +core/SynthesisExpr.v +core/SynthesisStmt.v +lib/LangDef.v +lib/IndexedMaps.v +core/SynthesisFunc.v +lib/Monad/StateMonadLaws.v +lib/Monad/StateMonad.v +lib/Monad/StateMonadOption.v +backend/Environments/Globalenvs.v +backend/Expressions/ExpMiniC.v +backend/Statements/StmtMiniC.v +core/Syntax.v +core/HyperTypeInst.v +lib/SimpleIndexedMaps.v +lib/SimpleMaps.v +lib/Monad/ContOptionMonad.v +lib/OProp.v +lib/Lens.v +core/MemoryModel.v +backend/AbstractData.v +backend/MachineModel.v +core/HyperType.v +backend/MemoryModel.v +backend/Environments/AllEnvironments.v +core/SEnv.v +lib/ArithInv.v +backend/Environments/ExtEnv.v +core/Cval.v +backend/SymbolicKeccak.v +lib/Monad/Monad.v +backend/Options.v +backend/Ctypes.v +backend/Values/HighValues.v +backend/IndexLib.v +backend/Values/LowValues.v +backend/Cop.v +Runtime.v diff --git a/src/backend/AST.v b/src/backend/AST.v new file mode 100755 index 0000000..541b0c1 --- /dev/null +++ b/src/backend/AST.v @@ -0,0 +1,36 @@ +(* This file is unchanged from CompCert except commenting things out. *) + + +(** This file defines a number of data types and operations used in + the abstract syntax trees of many of the intermediate languages. *) + +Require Import cclib.Coqlib. +Require String. +Require Import cclib.Integers. + +Set Implicit Arguments. + +(** * Syntactic elements *) + +(** Identifiers (names of local variables, of global symbols and functions, + etc) are represented by the type [positive] of positive integers. *) + +(* These are used. *) +Definition ident := positive. +Definition ident_eq := peq. + +Parameter ident_of_string : String.string -> ident. + +Definition label := ident. + +(* +(** Function definitions are the union of internal and external functions. +but we model external functions differently, so we don't need fundef *) +Definition fundef (F: Type): Type := F. +*) + + + (* this is not used: *) +(* | External: external_function -> fundef F *) + + diff --git a/src/backend/AbstractData.v b/src/backend/AbstractData.v new file mode 100644 index 0000000..7f5b766 --- /dev/null +++ b/src/backend/AbstractData.v @@ -0,0 +1,46 @@ +(* This code is adapted from liblayers.compcertx.AbstractData and + liblayers.compcertx.MemWithData. + + We should revisit it when we port liblayers to work with + DeepSEA/ethereum, but it shold be enough as a placeholder for + now. *) + +Require Import ZArith. +Require Import backend.Values.LowValues. +Require Import backend.MemoryModel. +Require Import backend.Environments.ExtEnv. + +(* In CompcertX, the type of memory is abstract and there is a + typeclass MemoryModelOps with operations on it, but for our + purposes we can just define it concretely. *) +Definition mem := ext_env. + + +Class CompatDataOps data := + { + empty_data : data; + high_level_invariant: data -> Prop + }. + +Class CompatData data `{data_ops: CompatDataOps data} := + { + empty_data_high_level_invariant: + high_level_invariant empty_data + }. + +Record compatdata := + { + cdata_type : Type; + cdata_ops : CompatDataOps cdata_type; + cdata_prf : CompatData cdata_type + }. + +Global Existing Instance cdata_ops. +Global Existing Instance cdata_prf. + +Definition cdata := Build_compatdata. +Global Arguments cdata _ {_ _}. + +Definition mwd (D: compatdata) := (mem * cdata_type D)%type. + + diff --git a/src/backend/Cop.v b/src/backend/Cop.v new file mode 100755 index 0000000..fa5d371 --- /dev/null +++ b/src/backend/Cop.v @@ -0,0 +1,151 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Arithmetic and logical operators *) + + +Require Import cclib.Coqlib. +Require Import backend.AST. +Require Import cclib.Integers. +(* Require Import Floats. *) +Require Import backend.Values.LowValues. +Require Import backend.Values.HighValues. +(*Require Import Memory. *) +Require Import backend.Ctypes. +Require Import SymbolicKeccak. + +(** * Syntax of operators. *) + +Inductive unary_operation : Type := + | Onotbool : unary_operation (**r boolean negation ([!] in C) *) + | Onotint : unary_operation (**r integer complement ([~] in C) *) + | Oneg : unary_operation (**r opposite (unary [-]) *) + | Osha_1 : unary_operation. (**r Keccak-256 hash. *) + +Inductive binary_operation : Type := + | Oadd : binary_operation (**r addition (binary [+]) *) + | Osub : binary_operation (**r subtraction (binary [-]) *) + | Omul : binary_operation (**r multiplication (binary [*]) *) + | Odiv : binary_operation (**r division ([/]) *) + | Omod : binary_operation (**r remainder ([%]) *) + | Oexp : binary_operation (**r exponentiation ([**]) *) + + | Oand : binary_operation (**r bitwise and ([&]) *) + | Oor : binary_operation (**r bitwise or ([|]) *) + | Oxor : binary_operation (**r bitwise xor ([^]) *) + | Oshl : binary_operation (**r left shift ([<<]) *) + | Oshr : binary_operation (**r right shift ([>>]) *) + | Oeq: binary_operation (**r comparison ([==]) *) + | One: binary_operation (**r comparison ([!=]) *) + | Olt: binary_operation (**r comparison ([<]) *) + | Ogt: binary_operation (**r comparison ([>]) *) + | Ole: binary_operation (**r comparison ([<=]) *) + | Oge: binary_operation (**r comparison ([>=]) *) + + | Osha_2 : binary_operation. (**r Keccak-256 hash *) + +Inductive incr_or_decr : Type := Incr | Decr. + + +Definition sem_unary_operation_int256 (op: unary_operation) (n: int256) : option int256 := + match op with + | Onotbool => Some (if (Int256.eq n Int256.zero) then Int256.one else Int256.zero) + | Onotint => Some (Int256.not n) + | Oneg => Some (Int256.neg n) + | Osha_1 => None (* handled elsewhere. *) + end. + +Definition sem_unary_operation (op: unary_operation) (v: val) (_:type) : option val := + match op with + | Osha_1 => Some (sha_1 v) + | _ => + match v with + | Vint i => match sem_unary_operation_int256 op i with + | Some i' => Some (Vint i') + | None => None + end + | _ => None + end + end. + +Definition int256_of_bool (b:bool) := + if b then Int256.one else Int256.zero. + +Definition Int256_comparison (signed:bool) (op : comparison) (n m : int256) := +if signed then Some (int256_of_bool (Int256.cmp op n m)) + else Some (int256_of_bool (Int256.cmpu op n m)). + +(* TODO: complete these. Need to have similar for all the other binary operations. *) +Definition sem_binary_operation_int256 (op: binary_operation) (signed:bool) (n m: int256) : option int256 := + match op with + | Oadd => Some (Int256.add n m) + | Osub => Some (Int256.sub n m) + | Omul => Some (Int256.mul n m) + | Odiv => Some (Int256.divu n m) + | Omod => if signed then Some (Int256.mods n m) + else Some (Int256.modu n m) + | Oexp => None (* todo *) + | Oand => Some (Int256.and n m) + | Oor => Some (Int256.or n m) + | Oxor => Some (Int256.xor n m) + | Oshl => Some (Int256.shl n m) + | Oshr => if signed then Some (Int256.shr n m) + else Some (Int256.shru n m) + | Oeq => Int256_comparison signed Ceq n m + | One => Int256_comparison signed Cne n m + | Olt => Int256_comparison signed Clt n m + | Ogt => Int256_comparison signed Cgt n m + | Oge => Int256_comparison signed Cge n m + | Ole => Int256_comparison signed Cle n m + + | Osha_2 => None (* handled elsewhere. *) + end. + +Definition val_eq_dec : forall (v1 v2 :val), {v1=v2}+{v1<>v2}. + decide equality. + + apply Int256.eq_dec. + + decide equality ; try apply Pos.eq_dec; try apply Int256.eq_dec; try apply IdentExtIndexed.eq. +Defined. + +Definition sem_binary_operation (op: binary_operation) (v1 :val) (t1:type) (v2 : val) (t2:type) : option val := + match op with + | Osha_2 => Some (sha_2 v1 v2) + | Oeq => Some (Vint (if val_eq_dec v1 v2 then (Int256.one) else (Int256.zero))) + | One => Some (Vint (if val_eq_dec v1 v2 then (Int256.zero) else (Int256.one))) + | _ => + match is_signed t1, is_signed t2 with + | false, false => + match v1, v2 with + | Vint i1, Vint i2 => match sem_binary_operation_int256 op false i1 i2 with + | Some i => Some (Vint i) + | None => None + end + | _,_ => None + end + | _,_ => None (* todo: support signed integer types *) + end + end. + +Definition sem_unop_low (op: unary_operation) (v:LowValues.val) (t:type) : option LowValues.val := + match sem_unary_operation op (OfLow v) t with + | Some v' => ToLowErr v' + | None => None + end. + +Definition sem_binop_low (op:binary_operation) (v1:LowValues.val) (t1:type) (v2:LowValues.val) (t2:type) := + match sem_binary_operation op (OfLow v1) t1 (OfLow v2) t2 with + | Some v' => ToLowErr v' + | None => None + end. diff --git a/src/backend/Ctypes.v b/src/backend/Ctypes.v new file mode 100755 index 0000000..4ea917f --- /dev/null +++ b/src/backend/Ctypes.v @@ -0,0 +1,154 @@ +(* These are the types we used, but the only ones we pay attention to are void, array, and struct. + Anything else is all treated the same. + +*) + +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Type expressions for the Compcert C and Clight languages *) + +Require Import cclib.Coqlib. +Require Import backend.AST. +Require Import cclib.Errors. +Require Import backend.Options. +(* Require Archi. *) + +(** * Syntax of types *) + +(** Compcert C types are similar to those of C. They include numeric types, + pointers, arrays, function types, and composite types (struct and + union). Numeric types (integers and floats) fully specify the + bit size of the type. An integer type is a pair of a signed/unsigned + flag and a bit size: 8, 16, or 32 bits, or the special [IBool] size + standing for the C99 [_Bool] type. 64-bit integers are treated separately. *) + +Inductive signedness : Type := + | Signed: signedness + | Unsigned: signedness. + +Inductive intsize : Type := + | I8: intsize + | I16: intsize + | I32: intsize + | I256: intsize + | IBool: intsize. + +Inductive ptrkind : Type := +| mem : ptrkind +| stor : ptrkind +| call : ptrkind. + +Inductive type : Type := + | Tvoid: type (**r the [void] type *) + | Tint: intsize -> signedness -> type (**r integer types *) + | Tpointer: ptrkind -> type -> type (**r pointer types ([*ty]) *) + | Tarray: type -> Z -> type (**r array types ([ty[len]]) *) + | Thashmap : type -> type -> type (** key_type, elem_type *) + | Tfunction: typelist -> type ->type (**r function types *) + | Tstruct: ident -> fieldlist -> type (**r struct types *) + | Tunion: ident -> fieldlist -> type (**r union types *) + | Tcomp_ptr: ident -> type (**r pointer to named struct or union *) + +with typelist : Type := + | Tnil: typelist + | Tcons: type -> typelist -> typelist + +with fieldlist : Type := + | Fnil: fieldlist + | Fcons: ident -> type -> fieldlist -> fieldlist. + +Lemma type_eq: forall (ty1 ty2: type), {ty1=ty2} + {ty1<>ty2} +with typelist_eq: forall (tyl1 tyl2: typelist), {tyl1=tyl2} + {tyl1<>tyl2} +with fieldlist_eq: forall (fld1 fld2: fieldlist), {fld1=fld2} + {fld1<>fld2}. +Proof. + assert (forall (x y: intsize), {x=y} + {x<>y}) by decide equality. + assert (forall (x y: signedness), {x=y} + {x<>y}) by decide equality. + generalize ident_eq zeq bool_dec. intros E1 E2 E3. + decide equality. + decide equality. + decide equality. + generalize ident_eq. intros E1. decide equality. +Defined. + +Opaque type_eq typelist_eq fieldlist_eq. + +Definition type_int32s := Tint I32 Signed. +Definition type_int256u := Tint I256 Unsigned. +Definition type_bool := Tint IBool Signed. + +(* Signedness. *) + +Definition is_signed (t: type) : bool := + match t with + | Tint _ Signed => true + | _ => false + end. + +(* Concatenation and length for fieldlsit. *) + +Fixpoint fieldlist_app (fld1 fld2: fieldlist) : fieldlist := + match fld1 with + | Fnil => fld2 + | Fcons id ty fld => Fcons id ty (fieldlist_app fld fld2) + end. + +Fixpoint fieldlist_length (fld: fieldlist) : nat := + match fld with + | Fnil => 0 + | Fcons id ty fld => S (fieldlist_length fld) + end. + +Lemma fieldlist_app_assoc fld1 fld2 fld3 : + fieldlist_app fld1 (fieldlist_app fld2 fld3) + = fieldlist_app (fieldlist_app fld1 fld2) fld3. +Proof. + induction fld1; [ reflexivity | simpl ]. + f_equal; apply IHfld1. +Qed. + +Lemma fieldlist_length_app fld1 fld2 : + fieldlist_length (fieldlist_app fld1 fld2) = + (fieldlist_length fld1 + fieldlist_length fld2)%nat. +Proof. + induction fld1; auto. + simpl. rewrite IHfld1. auto. +Qed. + +(** Size of a type in words *) +(* This is added to the EVM backend, which counts words sometimes. +Actually this should never matter, because everything is hashed *) + +Fixpoint sizeof_words (t: type) : Z := + match t with + | Tarray t' n => (sizeof_words t') * (Zmax 0 n) + | Tstruct _ fld => sizeof_struct_words fld 0 + | _ => 1 (* most things are one word *) + end + +with sizeof_struct_words (fld: fieldlist) (pos: Z) {struct fld} : Z := + match fld with + | Fnil => pos + | Fcons _ t fld' => sizeof_struct_words fld' (pos + (sizeof_words t)) + end. + +(* get offset and type *) +Fixpoint struct_field (fld: fieldlist) (id: ident) : option (nat * type) := + match fld with + | Fnil => None + | Fcons id' t fld' => if ident_eq id' id then Some (O, t) else + unpack offset, t' <- struct_field fld' id ;; + Some (S offset, t') + end. diff --git a/src/backend/Environments/AllEnvironments.v b/src/backend/Environments/AllEnvironments.v new file mode 100644 index 0000000..69c4b4b --- /dev/null +++ b/src/backend/Environments/AllEnvironments.v @@ -0,0 +1,41 @@ +Require Import cclib.Coqlib. +Require Import backend.Ctypes. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.Values.LowValues. +Require Import backend.Values.HighValues. + +(** These should be divided up into their own files. *) + +(* temp_env represents a local enviroment. It contains arguments and temps for a function. *) + +(* temps all have value 0 until initialized, except args *) +Definition temp_env : Type := PTree.t val. + +Definition temp_env_low : Type := PTree.t LowValues.val. + +(* stores contents of evm stack, which is is reality a list of values +but may conceptually be function args, temps during execution, or return values. + +The operational semantics use this (only a single stack_data, since the indidual step rules only case about the top part of the stack). +The highest language just uses Temps; then further compilation phases introduce Args, Retval etc. +*) +Inductive stack_data: Type := +| Temps: temp_env -> stack_data +| Args: list val -> stack_data (* When you have just called a function but not yet allocated temps. *) + +(* A function has n arguments and m temp variables. + The calling function pushes arguments. The function prelude pushes m zeros for the temps. + Then semantically, the rest of the function body is executed in an temp_env which has (n+m) entries. + There is an extra complication that if there are two function arguments with the same name, the first one will be shadowed and not put in the temp_env. + *) + +| MethodArgs: list val -> stack_data +| Retval: val -> stack_data. + +Inductive stack_data_low : Type := +| TempsL: temp_env_low -> stack_data_low +| ArgsL: list LowValues.val -> stack_data_low +| MethodArgsL: list LowValues.val -> stack_data_low +| RetvalL: LowValues.val -> stack_data_low +. diff --git a/src/backend/Environments/ExtEnv.v b/src/backend/Environments/ExtEnv.v new file mode 100644 index 0000000..87ece85 --- /dev/null +++ b/src/backend/Environments/ExtEnv.v @@ -0,0 +1,149 @@ +Require Import cclib.Coqlib. +Require Import backend.Ctypes. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.Values.HighValues. +Require Import backend.AST. + +Section ExtEnv. + +Variable function : Type. +Variable fn_locals : function -> list (ident * type). + +Inductive contents := +| EEVal : forall v:val, contents +| EEUndef : contents +| EEDead : contents +. + +(* at a lower level, storage and memory are separate *) +Definition ext_env : Type := IdentExtMap.t (contents). + +Notation "env [ i ]" := (IdentExtMap.get i env) (at level 80). +Notation "env [ i |-> v ]" := (IdentExtMap.set i v env) (at level 80). + +Definition empty_stack : ext_env := + Int256Map.init EEDead. + +Fixpoint alloc_eid (ty:type) (root:ident_ext) (ee:ext_env) (init:contents) : option ext_env := + match ty with + | Tarray ty' size => + let fix alloc (ind:nat) (ee:ext_env) := + match ind with + | O => Some ee + | S ind' => + match alloc ind' ee with + | Some ee' => alloc_eid ty' (Index root (Int256.repr (Z.of_nat ind))) ee' init + | None => None + end + end + in + match size with + | Zpos p => alloc (Pos.to_nat p) ee + | _ => None + end + | Tstruct id fl => + let fix alloc (fl:fieldlist) (ee:ext_env) := + match fl with + | Fnil => Some ee + | Fcons id ty' fl' => + match alloc fl' ee with + | Some se' => alloc_eid ty' (Field root id) se' init + | None => None + end + end + in alloc fl ee + | _ => Some (ee[root|->init]) + end. + +Fixpoint initial_locs (ctx:int256) (locs : list (ident*type)) (ee:ext_env) : option ext_env := + match locs with + | nil => Some ee + | (id,ty)::locs' => + match initial_locs ctx locs' ee with + | Some se' => alloc_eid ty (Field (Local ctx) id) ee EEUndef + | None => None + end + end. + +Fixpoint clear_locs (ctx:int256) (locs : list (ident*type)) (ee:ext_env) : option ext_env := + match locs with + | nil => Some ee + | (id,ty)::locs' => + match initial_locs ctx locs' ee with + | Some se' => alloc_eid ty (Field (Local ctx) id) ee EEDead + | None => None + end + end. + +Fixpoint frame_size (locs:list (ident * type)) : int256 := + match locs with + | nil => Int256.repr 0 + | cons (_,ty) locs' => Int256.add (Int256.repr (sizeof_words ty)) (frame_size locs') + end. + +Definition push_func (ctx:int256) (f:function) (ee:ext_env) : option (ext_env*int256) := + match initial_locs ctx (fn_locals f) ee with + | Some ee' => Some (ee', Int256.add ctx (frame_size (fn_locals f))) + | None => None + end. + +Definition pop_func (ctx:int256) (f:function) (ee:ext_env) : option (ext_env*int256) := + match clear_locs ctx (fn_locals f) ee with + | Some ee' => Some (ee', Int256.sub ctx (frame_size (fn_locals f))) + | None => None + end. + +Definition read (eid:ident_ext) (ee:ext_env) : option val := + match ee[eid] with + | EEVal v => Some v + | _ => None + end. + +Definition write (eid:ident_ext) (v:val) (ee:ext_env) : option ext_env := + match ee[eid] with + | EEDead => None + | _ => Some (ee[eid|->(EEVal v)]) + end. + +Definition var (ctx:int256) (id:ident) (ee:ext_env) : option ident_ext := + let eid := Field (Local ctx) id in + match ee[eid] with + | EEDead => None + | _ => Some eid + end. + +Definition glob (id:ident) (ee:ext_env) : option ident_ext := + Some (Field Global id). + +Definition addr (eid:ident_ext) : val := + Vptr (LVeid eid). + +Definition deref (v:val) : option ident_ext := + match v with + | Vptr (LVeid eid) => Some eid + | _ => None + end. + +Definition access_field (base:ident_ext) (ty:type) (fid:ident) : option ident_ext := + match ty with + | Tstruct _ fl => + match struct_field fl fid with + | Some _ => Some (Field base fid) + | None => None + end + | _ => None + end. + +Definition index_array (base:ident_ext) (ty:type) (ind:val) : option ident_ext := + match ind,ty with + | Vint off, Tarray ty' size => + if Int256.lt off (Int256.repr size) then + Some (Index base off) + else None + | Vint off, Thashmap _ _ => + Some (Index base off) + | _, _ => None + end. + +End ExtEnv. diff --git a/src/backend/Environments/Globalenvs.v b/src/backend/Environments/Globalenvs.v new file mode 100755 index 0000000..42a4b8b --- /dev/null +++ b/src/backend/Environments/Globalenvs.v @@ -0,0 +1,2357 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Global environments are a component of the dynamic semantics of + all languages involved in the compiler. A global environment + maps symbol names (names of functions and of global variables) + to the corresponding memory addresses. It also maps memory addresses + of functions to the corresponding function descriptions. + + Global environments, along with the initial memory state at the beginning + of program execution, are built from the program of interest, as follows: +- A distinct memory address is assigned to each function of the program. + These function addresses use negative numbers to distinguish them from + addresses of memory blocks. The associations of function name to function + address and function address to function description are recorded in + the global environment. +- For each global variable, a memory block is allocated and associated to + the name of the variable. + + These operations reflect (at a high level of abstraction) what takes + place during program linking and program loading in a real operating + system. *) + +Require Import cclib.Coqlib. +Require Import cclib.Errors. +Require Import cclib.Maps. +Require Import backend.AST. +Require Import cclib.Integers. +Require Import backend.Values.LowValues. +Require Import backend.Options. + +(* + +(** * Symbol environments *) + +(** Symbol environments are a restricted view of global environments, + focusing on symbol names and their associated blocks. They do not + contain mappings from blocks to function or variable definitions. *) + +Module Senv. + +Record t: Type := mksenv { + (** Operations *) + find_symbol: ident -> option block; + public_symbol: ident -> bool; + invert_symbol: block -> option ident; + block_is_volatile: block -> option bool; + nextblock: block; + (** Properties *) + find_symbol_injective: + forall id1 id2 b, find_symbol id1 = Some b -> find_symbol id2 = Some b -> id1 = id2; + invert_find_symbol: + forall id b, invert_symbol b = Some id -> find_symbol id = Some b; + find_invert_symbol: + forall id b, find_symbol id = Some b -> invert_symbol b = Some id; + public_symbol_exists: + forall id, public_symbol id = true -> exists b, find_symbol id = Some b; + find_symbol_below: + forall id b, find_symbol id = Some b -> Plt b nextblock; + block_is_volatile_below: + forall b, block_is_volatile b = Some true -> Plt b nextblock +}. + +Definition symbol_address (ge: t) (id: ident) (ofs: ptrofs) : val := + match find_symbol ge id with + | Some b => Vptr b ofs + | None => Vundef + end. + +Theorem shift_symbol_address: + forall ge id ofs delta, + symbol_address ge id (Ptrofs.add ofs delta) = Val.offset_ptr (symbol_address ge id ofs) delta. +Proof. + intros. unfold symbol_address, Val.offset_ptr. destruct (find_symbol ge id); auto. +Qed. + +Theorem shift_symbol_address_32: + forall ge id ofs n, + Archi.ptr64 = false -> + symbol_address ge id (Ptrofs.add ofs (Ptrofs.of_int n)) = Val.add (symbol_address ge id ofs) (Vint n). +Proof. + intros. unfold symbol_address. destruct (find_symbol ge id). +- unfold Val.add. rewrite H. auto. +- auto. +Qed. + +Theorem shift_symbol_address_64: + forall ge id ofs n, + Archi.ptr64 = true -> + symbol_address ge id (Ptrofs.add ofs (Ptrofs.of_int64 n)) = Val.addl (symbol_address ge id ofs) (Vlong n). +Proof. + intros. unfold symbol_address. destruct (find_symbol ge id). +- unfold Val.addl. rewrite H. auto. +- auto. +Qed. + +Definition equiv (se1 se2: t) : Prop := + nextblock se2 = nextblock se1 + /\ + (forall id, find_symbol se2 id = find_symbol se1 id) + /\ (forall id, public_symbol se2 id = public_symbol se1 id) + /\ (forall b, block_is_volatile se2 b = block_is_volatile se1 b). + +End Senv. +*) + + +Module IntIndexed <: INDEXED_TYPE. + +Definition t := int. +Definition index (i : int) := + let (intval, _) := i in + match intval with + | Z0 => xH + | Zpos p => xO p + | Zneg p => xI p + end. + +Definition index_inv p := + match p with + | xH => Int.zero + | xO p' => Int.repr (Zpos p') + | xI p' => Int.repr (Zneg p') + end. + + Lemma zlt_proof_irrelevance : + forall {x y : Z} (p q : x < y), p = q. + Proof. + intros. + apply Eqdep_dec.eq_proofs_unicity. + decide equality. + Qed. + + Lemma range_proof_irrelevance : + forall {x y z : Z} (p q : x < y < z), p = q. + Proof. + intros. + destruct p as [p1 p2]. + destruct q as [q1 q2]. + rewrite (zlt_proof_irrelevance p1 q1). + rewrite (zlt_proof_irrelevance p2 q2). + reflexivity. + Qed. + + Transparent Int.Z_mod_modulus. + Lemma modulus_nop : forall x, + -1 < x < Int.modulus -> Int.Z_mod_modulus x = x. + Proof. + intros. + unfold Int.Z_mod_modulus. + destruct x. + + reflexivity. + + rewrite Int.P_mod_two_p_eq. + rewrite Zmod_small by (unfold Int.modulus in H; omega). + reflexivity. + + destruct H. + unfold Z.lt in H. + simpl in H. + destruct p; simpl in H; congruence. + Qed. + Opaque Int.Z_mod_modulus. + + Transparent Int.repr. + Lemma index_invertible : forall x, index_inv (index x) = x. + Proof. + intros. + destruct x as [x x_range]. + destruct x; simpl. + + unfold Int.zero. + unfold Int.repr. + rewrite (range_proof_irrelevance (Int.Z_mod_modulus_range' 0) x_range). + reflexivity. + + unfold Int.repr. + generalize (Int.Z_mod_modulus_range' (Z.pos p)). + rewrite (modulus_nop _ x_range). + intros a. + rewrite (range_proof_irrelevance a x_range). + reflexivity. + + unfold Int.repr. + generalize (Int.Z_mod_modulus_range' (Z.neg p)). + rewrite (modulus_nop _ x_range). + intros a. + rewrite (range_proof_irrelevance a x_range). + reflexivity. + Qed. + Opaque Int.repr. + + Lemma index_inj: forall (x y: int), index x = index y -> x = y. + Proof. + unfold index; destruct x; destruct y; intros; + try discriminate; try reflexivity. + apply Int.mkint_eq. + destruct intval; destruct intval0; try (inversion H); try reflexivity. + Qed. + + Definition eq := Int.eq_dec. + +End IntIndexed. + +Module IntMap := IMap(IntIndexed). + + + +Module Genv. + +(** * Global environments *) + +Section GENV. + +Variable F: Type. (**r The type of function descriptions *) +Variable V: Type. (**r The type of information attached to variables *) + +(** The type of global environments. *) + + +Record t: Type := mkgenv { + (* list of variables *) + genv_vars: list ident; (**r which symbol names are public *) + (* list of functions *) + genv_funcs: list ident; (**r which symbol names are public *) + (* list of methods, i.e. public functions *) + genv_methods: list int; + (* type of variables *) + genv_defs: PTree.t V; (**r mapping symbol -> type *) + (* function definitions *) + genv_fundefs: PTree.t F; (**r mapping symbol -> function definition *) + (* method definitions *) + genv_methoddefs: IntMap.t (option F); + (* constructor *) + genv_constructor: option F; +}. + +(* + +Record t: Type := mkgenv { + (* list of variables *) + genv_public: list ident; (**r which symbol names are public *) + (* where in memory *) + genv_symb: PTree.t block; (**r mapping symbol -> block *) + (* type of variables *) + genv_defs: PTree.t V; (**r mapping block -> definition *) + genv_fundefs: PTree.t F; (**r mapping block -> definition *) + genv_next: block; (**r next symbol pointer *) + genv_symb_range: forall id b, PTree.get id genv_symb = Some b -> Plt b genv_next; + genv_defs_range: forall b g, PTree.get b genv_defs = Some g -> Plt b genv_next; + genv_vars_inj: forall id1 id2 b, + PTree.get id1 genv_symb = Some b -> PTree.get id2 genv_symb = Some b -> id1 = id2 +}. + +*) + +(** Creation functions *) + +Definition empty_genv : t := + mkgenv nil nil nil (PTree.empty _) (PTree.empty _) (IntMap.init None) None. + +Fixpoint add_genv_funcs (funcs: list (ident * F)) (ge: t) : t := + match funcs with + | nil => ge + | (id, fundef)::rest => let r := add_genv_funcs rest ge in + mkgenv r.(genv_vars) (id::(r.(genv_funcs))) r.(genv_methods) r.(genv_defs) (PTree.set id fundef r.(genv_fundefs)) r.(genv_methoddefs) r.(genv_constructor) + end. + +Fixpoint add_genv_vars (vars: list (ident * V)) (ge: t) : t := + match vars with + | nil => ge + | (id, def)::rest => let r := add_genv_vars rest ge in + mkgenv (id::(r.(genv_vars))) r.(genv_funcs) r.(genv_methods) (PTree.set id def r.(genv_defs)) r.(genv_fundefs) r.(genv_methoddefs) r.(genv_constructor) + end. + +Fixpoint add_genv_methods (methods: list (int * F)) (ge: t) : t := + match methods with + | nil => ge + | (sig, fundef)::rest => let r := add_genv_methods rest ge in + mkgenv r.(genv_vars) r.(genv_funcs) (sig::(r.(genv_methods))) r.(genv_defs) r.(genv_fundefs) (IntMap.set sig (Some fundef) r.(genv_methoddefs)) r.(genv_constructor) + end. + +Definition add_genv_constructor (constructor: option F) (ge: t) : t := + mkgenv ge.(genv_vars) ge.(genv_funcs) ge.(genv_methods) ge.(genv_defs) ge.(genv_fundefs) ge.(genv_methoddefs) constructor. + +Definition new_genv (vars: list (ident * V)) + (funcs: list (ident * F)) (methods: list (int * F)) + (constructor: option F) : t := + add_genv_constructor constructor (add_genv_methods methods (add_genv_funcs funcs (add_genv_vars vars empty_genv))). + +(* This includes functions and methods, but not the constructor. The difference is that you are not allowed to + call the constructor. *) +Definition all_functions (ge: t) : list F := + (map snd (PTree.elements ge.(genv_fundefs))) + ++ (flat_map (optional_filter F) (map snd (PTree.elements (snd ge.(genv_methoddefs))))) + . + +(* (map snd (PTree.elements ge.(genv_fundefs))) + ++ (flat_map (optional_filter F) (map snd (PTree.elements (snd ge.(genv_methoddefs))))) + ++ (optional_filter F ge.(genv_constructor)).*) + +Definition count_methods (ge: t) : nat := + length (PTree.elements (snd (genv_methoddefs ge))). + +(* + +(** [find_symbol ge id] returns the block associated with the given name, if any *) + +Definition find_symbol (ge: t) (id: ident) : option block := + PTree.get id ge.(genv_symb). + +(** [symbol_address ge id ofs] returns a pointer into the block associated + with [id], at byte offset [ofs]. [Vundef] is returned if no block is associated + to [id]. *) + +Definition symbol_address (ge: t) (id: ident) (ofs: ptrofs) : val := + match find_symbol ge id with + | Some b => Vptr b ofs + | None => Vundef + end. + +(** [public_symbol ge id] says whether the name [id] is public and defined. *) + +Definition public_symbol (ge: t) (id: ident) : bool := + match find_symbol ge id with + | None => false + | Some _ => In_dec ident_eq id ge.(genv_public) + end. + +(** [find_def ge b] returns the global definition associated with the given address. *) + +Definition find_def (ge: t) (b: block) : option (globdef F V) := + PTree.get b ge.(genv_defs). + +(** [find_funct_ptr ge b] returns the function description associated with + the given address. *) + +Definition find_funct_ptr (ge: t) (b: block) : option F := + match find_def ge b with Some (Gfun f) => Some f | _ => None end. + +(** [find_funct] is similar to [find_funct_ptr], but the function address + is given as a value, which must be a pointer with offset 0. *) + +Definition find_funct (ge: t) (v: val) : option F := + match v with + | Vptr b ofs => if Ptrofs.eq_dec ofs Ptrofs.zero then find_funct_ptr ge b else None + | _ => None + end. + +(** [invert_symbol ge b] returns the name associated with the given block, if any *) + +Definition invert_symbol (ge: t) (b: block) : option ident := + PTree.fold + (fun res id b' => if eq_block b b' then Some id else res) + ge.(genv_symb) None. + +(** [find_var_info ge b] returns the information attached to the variable + at address [b]. *) + +Definition find_var_info (ge: t) (b: block) : option (globvar V) := + match find_def ge b with Some (Gvar v) => Some v | _ => None end. + +(** [block_is_volatile ge b] returns [true] if [b] points to a global variable + of volatile type, [false] otherwise. *) + +Definition block_is_volatile (ge: t) (b: block) : option bool := + match find_var_info ge b with + | None => None + | Some gv => Some (gv.(gvar_volatile)) + end. + +(** ** Constructing the global environment *) + +Program Definition add_global (ge: t) (idg: ident * option (globdef F V)) : t := + @mkgenv + ge.(genv_public) + (PTree.set idg#1 ge.(genv_next) ge.(genv_symb)) + (match idg#2 with Some g => PTree.set ge.(genv_next) g ge.(genv_defs) | _ => ge.(genv_defs) end) + (Psucc ge.(genv_next)) + _ _ _. +Next Obligation. + destruct ge; simpl in *. + rewrite PTree.gsspec in H. destruct (peq id i). inv H. apply Plt_succ. + apply Plt_trans_succ; eauto. +Qed. +Next Obligation. + destruct ge; simpl in *. + destruct o. + + + rewrite PTree.gsspec in H. destruct (peq b genv_next0). + inv H. apply Plt_succ. + apply Plt_trans_succ; eauto. + + + apply genv_defs_range0 in H. + xomega. +Qed. +Next Obligation. + destruct ge; simpl in *. + rewrite PTree.gsspec in H. rewrite PTree.gsspec in H0. + destruct (peq id1 i); destruct (peq id2 i). + congruence. + inv H. eelim Plt_strict. eapply genv_symb_range0; eauto. + inv H0. eelim Plt_strict. eapply genv_symb_range0; eauto. + eauto. +Qed. + +Definition add_globals (ge: t) (gl: list (ident * option (globdef F V))) : t := + List.fold_left add_global gl ge. + +Lemma add_globals_app: + forall gl2 gl1 ge, + add_globals ge (gl1 ++ gl2) = add_globals (add_globals ge gl1) gl2. +Proof. + intros. apply fold_left_app. +Qed. + +Program Definition empty_genv (pub: list ident): t := + @mkgenv pub (PTree.empty _) (PTree.empty _) 1%positive _ _ _. +Next Obligation. + rewrite PTree.gempty in H. discriminate. +Qed. +Next Obligation. + rewrite PTree.gempty in H. discriminate. +Qed. +Next Obligation. + rewrite PTree.gempty in H. discriminate. +Qed. + +Definition globalenv (p: program F V) := + add_globals (empty_genv p.(prog_public)) p.(prog_defs). + +(** Proof principles *) + +Section GLOBALENV_PRINCIPLES. + +Variable P: t -> Prop. + +Lemma add_globals_preserves: + forall gl ge, + (forall ge id g, P ge -> In (id, g) gl -> P (add_global ge (id, g))) -> + P ge -> P (add_globals ge gl). +Proof. + induction gl; simpl; intros. + auto. + destruct a. apply IHgl; auto. +Qed. + +Lemma add_globals_ensures: + forall id g gl ge, + (forall ge id g, P ge -> In (id, g) gl -> P (add_global ge (id, g))) -> + (forall ge, P (add_global ge (id, g))) -> + In (id, g) gl -> P (add_globals ge gl). +Proof. + induction gl; simpl; intros. + contradiction. + destruct H1. subst a. apply add_globals_preserves; auto. + apply IHgl; auto. +Qed. + +Lemma add_globals_unique_preserves: + forall id gl ge, + (forall ge id1 g, P ge -> In (id1, g) gl -> id1 <> id -> P (add_global ge (id1, g))) -> + ~In id (map fst gl) -> P ge -> P (add_globals ge gl). +Proof. + induction gl; simpl; intros. + auto. + destruct a. apply IHgl; auto. +Qed. + +Lemma add_globals_unique_ensures: + forall gl1 id g gl2 ge, + (forall ge id1 g1, P ge -> In (id1, g1) gl2 -> id1 <> id -> P (add_global ge (id1, g1))) -> + (forall ge, P (add_global ge (id, g))) -> + ~In id (map fst gl2) -> P (add_globals ge (gl1 ++ (id, g) :: gl2)). +Proof. + intros. rewrite add_globals_app. simpl. apply add_globals_unique_preserves with id; auto. +Qed. + +Remark in_norepet_unique: + forall (A: Type), + forall id g (gl: list (ident * A)), + In (id, g) gl -> list_norepet (map fst gl) -> + exists gl1 gl2, gl = gl1 ++ (id, g) :: gl2 /\ ~In id (map fst gl2). +Proof. + induction gl as [|[id1 g1] gl]; simpl; intros. + contradiction. + inv H0. destruct H. + inv H. exists nil, gl. auto. + exploit IHgl; eauto. intros (gl1 & gl2 & X & Y). + exists ((id1, g1) :: gl1), gl2; split; auto. rewrite X; auto. +Qed. + +Lemma add_globals_norepet_ensures: + forall id g gl ge, + (forall ge id1 g1, P ge -> In (id1, g1) gl -> id1 <> id -> P (add_global ge (id1, g1))) -> + (forall ge, P (add_global ge (id, g))) -> + In (id, g) gl -> list_norepet (map fst gl) -> P (add_globals ge gl). +Proof. + intros. exploit in_norepet_unique; eauto. intros (gl1 & gl2 & X & Y). + subst gl. apply add_globals_unique_ensures; auto. intros. eapply H; eauto. + apply in_or_app; simpl; auto. +Qed. + +End GLOBALENV_PRINCIPLES. + +(** ** Properties of the operations over global environments *) + +Theorem public_symbol_exists: + forall ge id, public_symbol ge id = true -> exists b, find_symbol ge id = Some b. +Proof. + unfold public_symbol; intros. destruct (find_symbol ge id) as [b|]. + exists b; auto. + discriminate. +Qed. + +Theorem shift_symbol_address: + forall ge id ofs delta, + symbol_address ge id (Ptrofs.add ofs delta) = Val.offset_ptr (symbol_address ge id ofs) delta. +Proof. + intros. unfold symbol_address, Val.offset_ptr. destruct (find_symbol ge id); auto. +Qed. + +Theorem shift_symbol_address_32: + forall ge id ofs n, + Archi.ptr64 = false -> + symbol_address ge id (Ptrofs.add ofs (Ptrofs.of_int n)) = Val.add (symbol_address ge id ofs) (Vint n). +Proof. + intros. unfold symbol_address. destruct (find_symbol ge id). +- unfold Val.add. rewrite H. auto. +- auto. +Qed. + +Theorem shift_symbol_address_64: + forall ge id ofs n, + Archi.ptr64 = true -> + symbol_address ge id (Ptrofs.add ofs (Ptrofs.of_int64 n)) = Val.addl (symbol_address ge id ofs) (Vlong n). +Proof. + intros. unfold symbol_address. destruct (find_symbol ge id). +- unfold Val.addl. rewrite H. auto. +- auto. +Qed. + +Theorem find_funct_inv: + forall ge v f, + find_funct ge v = Some f -> exists b, v = Vptr b Ptrofs.zero. +Proof. + intros until f; unfold find_funct. + destruct v; try congruence. + destruct (Ptrofs.eq_dec i Ptrofs.zero); try congruence. + intros. exists b; congruence. +Qed. + +Theorem find_funct_find_funct_ptr: + forall ge b, + find_funct ge (Vptr b Ptrofs.zero) = find_funct_ptr ge b. +Proof. + intros; simpl. apply dec_eq_true. +Qed. + +Theorem find_funct_ptr_iff: + forall ge b f, find_funct_ptr ge b = Some f <-> find_def ge b = Some (Gfun f). +Proof. + intros. unfold find_funct_ptr. destruct (find_def ge b) as [[f1|v1]|]; intuition congruence. +Qed. + +Theorem find_var_info_iff: + forall ge b v, find_var_info ge b = Some v <-> find_def ge b = Some (Gvar v). +Proof. + intros. unfold find_var_info. destruct (find_def ge b) as [[f1|v1]|]; intuition congruence. +Qed. + +Theorem find_def_symbol: + forall p id g, + (prog_defmap p)!id = Some g <-> exists b, find_symbol (globalenv p) id = Some b /\ find_def (globalenv p) b = Some g. +Proof. + intros. + set (P := fun m ge => m!id = Some g <-> exists b, find_symbol ge id = Some b /\ find_def ge b = Some g). + assert (REC: forall l m ge, + P m ge -> + P (fold_left + (fun m idg => + (*PTree.set idg#1 idg#2 m*) + match idg#2 with Some g => PTree.set idg#1 g m | _ => PTree.remove idg#1 m end + ) + l m) + (add_globals ge l)). + { induction l as [ | [id1 g1] l]; intros; simpl. + - auto. + - apply IHl. unfold P, add_global, find_symbol, find_def; simpl. + destruct g1. + * + rewrite ! PTree.gsspec. destruct (peq id id1). + + subst id1. split; intros. + inv H0. exists (genv_next ge); split; auto. apply PTree.gss. + destruct H0 as (b & A & B). inv A. rewrite PTree.gss in B. auto. + + red in H; rewrite H. split. + intros (b & A & B). exists b; split; auto. rewrite PTree.gso; auto. + apply Plt_ne. eapply genv_symb_range; eauto. + intros (b & A & B). rewrite PTree.gso in B. exists b; auto. + apply Plt_ne. eapply genv_symb_range; eauto. + * + rewrite PTree.gsspec. + rewrite PTree.grspec. + destruct (peq id id1); destruct (PTree.elt_eq id id1); auto; try intuition congruence. + subst id1. + split; try discriminate. + destruct 1 as (? & J & K). + inv J. + apply genv_defs_range in K. + xomega. + } + apply REC. unfold P, find_symbol, find_def; simpl. + rewrite ! PTree.gempty. split. + congruence. + intros (b & A & B); congruence. +Qed. + +Theorem find_def_symbol_strong: + forall p id g, + (prog_option_defmap p)!id = Some g <-> exists b, find_symbol (globalenv p) id = Some b /\ find_def (globalenv p) b = g. +Proof. + intros. + set (P := fun m ge => m!id = Some g <-> exists b, find_symbol ge id = Some b /\ find_def ge b = g). + assert (REC: forall l m ge, + P m ge -> + P (fold_left + (fun m idg => PTree.set idg#1 idg#2 m) + l m) + (add_globals ge l)). + { induction l as [ | [id1 g1] l]; intros; simpl. + - auto. + - apply IHl. unfold P, add_global, find_symbol, find_def; simpl. + rewrite ! PTree.gsspec. destruct (peq id id1). + + subst id1. split; intros. + inv H0. exists (genv_next ge); split; auto. + destruct g. + { apply PTree.gss. } + { destruct (_ ! _) eqn:GE; auto. + apply genv_defs_range in GE. + xomega. } + destruct H0 as (b & A & B). clear P IHl H. inv A. + destruct g1. + { rewrite PTree.gss. auto. } + destruct (_ ! _) eqn:GE; auto. + apply genv_defs_range in GE. + xomega. + + red in H; rewrite H. split. + intros (b & A & B). exists b; split; auto. + destruct g1; auto. + rewrite PTree.gso; auto. + apply Plt_ne. eapply genv_symb_range; eauto. + intros (b & A & B). + destruct g1; eauto. + rewrite PTree.gso in B. exists b; auto. + apply Plt_ne. eapply genv_symb_range; eauto. + } + apply REC. unfold P, find_symbol, find_def; simpl. + rewrite ! PTree.gempty. split. + congruence. + intros (b & A & B); congruence. +Qed. + +Theorem find_symbol_exists: + forall p id g, + In (id, g) (prog_defs p) -> + exists b, find_symbol (globalenv p) id = Some b. +Proof. + intros. unfold globalenv. eapply add_globals_ensures; eauto. +(* preserves *) + intros. unfold find_symbol; simpl. rewrite PTree.gsspec. destruct (peq id id0). + econstructor; eauto. + auto. +(* ensures *) + intros. unfold find_symbol; simpl. rewrite PTree.gss. econstructor; eauto. +Qed. + +Theorem find_symbol_inversion : forall p x b, + find_symbol (globalenv p) x = Some b -> + In x (prog_defs_names p). +Proof. + intros until b; unfold globalenv. eapply add_globals_preserves. +(* preserves *) + unfold find_symbol; simpl; intros. rewrite PTree.gsspec in H1. + destruct (peq x id). subst x. change id with (fst (id, g)). apply List.in_map; auto. + auto. +(* base *) + unfold find_symbol; simpl; intros. rewrite PTree.gempty in H. discriminate. +Qed. + +Theorem find_def_inversion: + forall p b g, + find_def (globalenv p) b = Some g -> + exists id, In (id, Some g) (prog_defs p). +Proof. + intros until g. unfold globalenv. apply add_globals_preserves. +(* preserves *) + unfold find_def; simpl; intros. + destruct g0; auto. + rewrite PTree.gsspec in H1. destruct (peq b (genv_next ge)). + inv H1. exists id; auto. + auto. +(* base *) + unfold find_def; simpl; intros. rewrite PTree.gempty in H. discriminate. +Qed. + +Corollary find_funct_ptr_inversion: + forall p b f, + find_funct_ptr (globalenv p) b = Some f -> + exists id, In (id, Some (Gfun f)) (prog_defs p). +Proof. + intros. apply find_def_inversion with b. apply find_funct_ptr_iff; auto. +Qed. + +Corollary find_funct_inversion: + forall p v f, + find_funct (globalenv p) v = Some f -> + exists id, In (id, Some (Gfun f)) (prog_defs p). +Proof. + intros. exploit find_funct_inv; eauto. intros [b EQ]. subst v. + rewrite find_funct_find_funct_ptr in H. + eapply find_funct_ptr_inversion; eauto. +Qed. + +Theorem find_funct_ptr_prop: + forall (P: F -> Prop) p b f, + (forall id f, In (id, Some (Gfun f)) (prog_defs p) -> P f) -> + find_funct_ptr (globalenv p) b = Some f -> + P f. +Proof. + intros. exploit find_funct_ptr_inversion; eauto. intros [id IN]. eauto. +Qed. + +Theorem find_funct_prop: + forall (P: F -> Prop) p v f, + (forall id f, In (id, Some (Gfun f)) (prog_defs p) -> P f) -> + find_funct (globalenv p) v = Some f -> + P f. +Proof. + intros. exploit find_funct_inversion; eauto. intros [id IN]. eauto. +Qed. + +Theorem global_addresses_distinct: + forall ge id1 id2 b1 b2, + id1 <> id2 -> + find_symbol ge id1 = Some b1 -> + find_symbol ge id2 = Some b2 -> + b1 <> b2. +Proof. + intros. red; intros; subst. elim H. destruct ge. eauto. +Qed. + +Theorem invert_find_symbol: + forall ge id b, + invert_symbol ge b = Some id -> find_symbol ge id = Some b. +Proof. + intros until b; unfold find_symbol, invert_symbol. + apply PTree_Properties.fold_rec. + intros. rewrite H in H0; auto. + congruence. + intros. destruct (eq_block b v). inv H2. apply PTree.gss. + rewrite PTree.gsspec. destruct (peq id k). + subst. assert (m!k = Some b) by auto. congruence. + auto. +Qed. + +Theorem find_invert_symbol: + forall ge id b, + find_symbol ge id = Some b -> invert_symbol ge b = Some id. +Proof. + intros until b. + assert (find_symbol ge id = Some b -> exists id', invert_symbol ge b = Some id'). + unfold find_symbol, invert_symbol. + apply PTree_Properties.fold_rec. + intros. rewrite H in H0; auto. + rewrite PTree.gempty; congruence. + intros. destruct (eq_block b v). exists k; auto. + rewrite PTree.gsspec in H2. destruct (peq id k). + inv H2. congruence. auto. + + intros; exploit H; eauto. intros [id' A]. + assert (id = id'). eapply genv_vars_inj; eauto. apply invert_find_symbol; auto. + congruence. +Qed. + +Definition advance_next (gl: list (ident * option (globdef F V))) (x: positive) := + List.fold_left (fun n g => Psucc n) gl x. + +Remark genv_next_add_globals: + forall gl ge, + genv_next (add_globals ge gl) = advance_next gl (genv_next ge). +Proof. + induction gl; simpl; intros. + auto. + rewrite IHgl. auto. +Qed. + +Remark genv_public_add_globals: + forall gl ge, + genv_public (add_globals ge gl) = genv_public ge. +Proof. + induction gl; simpl; intros. + auto. + rewrite IHgl; auto. +Qed. + +Theorem globalenv_public: + forall p, genv_public (globalenv p) = prog_public p. +Proof. + unfold globalenv; intros. rewrite genv_public_add_globals. auto. +Qed. + +Theorem block_is_volatile_below: + forall ge b, block_is_volatile ge b = Some true -> Plt b ge.(genv_next). +Proof. + unfold block_is_volatile; intros. destruct (find_var_info ge b) as [gv|] eqn:FV. + rewrite find_var_info_iff in FV. eapply genv_defs_range; eauto. + discriminate. +Qed. + +(** ** Coercing a global environment into a symbol environment *) + +Definition to_senv (ge: t) : Senv.t := + @Senv.mksenv + (find_symbol ge) + (public_symbol ge) + (invert_symbol ge) + (block_is_volatile ge) + ge.(genv_next) + ge.(genv_vars_inj) + (invert_find_symbol ge) + (find_invert_symbol ge) + (public_symbol_exists ge) + ge.(genv_symb_range) + (block_is_volatile_below ge). +*) +End GENV. +(* +(** * Construction of the initial memory state *) + +Section WITHMEMORYMODEL. +Context `{memory_model_prf: Mem.MemoryModel}. + +Variable F: Type. (**r The type of function descriptions *) +Variable V: Type. (**r The type of information attached to variables *) + +Section INITMEM. + +Variable ge: t F V. + +Definition store_init_data (m: mem) (b: block) (p: Z) (id: init_data) : option mem := + match id with + | Init_int8 n => Mem.store Mint8unsigned m b p (Vint n) + | Init_int16 n => Mem.store Mint16unsigned m b p (Vint n) + | Init_int32 n => Mem.store Mint32 m b p (Vint n) + | Init_int64 n => Mem.store Mint64 m b p (Vlong n) + | Init_float32 n => Mem.store Mfloat32 m b p (Vsingle n) + | Init_float64 n => Mem.store Mfloat64 m b p (Vfloat n) + | Init_addrof symb ofs => + match find_symbol ge symb with + | None => None + | Some b' => Mem.store Mptr m b p (Vptr b' ofs) + end + | Init_space n => Some m + end. + +Fixpoint store_init_data_list (m: mem) (b: block) (p: Z) (idl: list init_data) + {struct idl}: option mem := + match idl with + | nil => Some m + | id :: idl' => + match store_init_data m b p id with + | None => None + | Some m' => store_init_data_list m' b (p + init_data_size id) idl' + end + end. + +Definition perm_globvar (gv: globvar V) : permission := + if gv.(gvar_volatile) then Nonempty + else if gv.(gvar_readonly) then Readable + else Writable. + +Definition alloc_global (m: mem) (idg: ident * option (globdef F V)): option mem := + match idg with + | (id, None) => + let (m1, b) := Mem.alloc m 0 0 in + Some m1 + | (id, Some (Gfun f)) => + let (m1, b) := Mem.alloc m 0 1 in + Mem.drop_perm m1 b 0 1 Nonempty + | (id, Some (Gvar v)) => + let init := v.(gvar_init) in + let sz := init_data_list_size init in + let (m1, b) := Mem.alloc m 0 sz in + match store_zeros m1 b 0 sz with + | None => None + | Some m2 => + match store_init_data_list m2 b 0 init with + | None => None + | Some m3 => Mem.drop_perm m3 b 0 sz (perm_globvar v) + end + end + end. + +Fixpoint alloc_globals (m: mem) (gl: list (ident * option (globdef F V))) + {struct gl} : option mem := + match gl with + | nil => Some m + | g :: gl' => + match alloc_global m g with + | None => None + | Some m' => alloc_globals m' gl' + end + end. + +Lemma alloc_globals_app : forall gl1 gl2 m m1, + alloc_globals m gl1 = Some m1 -> + alloc_globals m1 gl2 = alloc_globals m (gl1 ++ gl2). +Proof. + induction gl1. + simpl. intros. inversion H; subst. auto. + simpl. intros. destruct (alloc_global m a); eauto. inversion H. +Qed. + +(** Next-block properties *) + +Remark store_zeros_nextblock: + forall m b p n m', store_zeros m b p n = Some m' -> Mem.nextblock m' = Mem.nextblock m. +Proof. + intros until n. functional induction (store_zeros m b p n); intros. + inv H; auto. + rewrite IHo; eauto with mem. + congruence. +Qed. + +Remark store_init_data_list_nextblock: + forall idl b m p m', + store_init_data_list m b p idl = Some m' -> + Mem.nextblock m' = Mem.nextblock m. +Proof. + induction idl; simpl; intros until m'. + intros. congruence. + caseEq (store_init_data m b p a); try congruence. intros. + transitivity (Mem.nextblock m0). eauto. + destruct a; simpl in H; try (eapply Mem.nextblock_store; eauto; fail). + congruence. + destruct (find_symbol ge i); try congruence. eapply Mem.nextblock_store; eauto. +Qed. + +Remark alloc_global_nextblock: + forall g m m', + alloc_global m g = Some m' -> + Mem.nextblock m' = Psucc(Mem.nextblock m). +Proof. + unfold alloc_global. intros. + destruct g as [id [[f|v]|]]. + (* function *) + destruct (Mem.alloc m 0 1) as [m1 b] eqn:?. + erewrite Mem.nextblock_drop; eauto. erewrite Mem.nextblock_alloc; eauto. + (* variable *) + set (init := gvar_init v) in *. + set (sz := init_data_list_size init) in *. + destruct (Mem.alloc m 0 sz) as [m1 b] eqn:?. + destruct (store_zeros m1 b 0 sz) as [m2|] eqn:?; try discriminate. + destruct (store_init_data_list m2 b 0 init) as [m3|] eqn:?; try discriminate. + erewrite Mem.nextblock_drop; eauto. + erewrite store_init_data_list_nextblock; eauto. + erewrite store_zeros_nextblock; eauto. + erewrite Mem.nextblock_alloc; eauto. + (* none *) + destruct (Mem.alloc m 0 0) as [m1 b] eqn:? . + inv H. + eapply Mem.nextblock_alloc. + eassumption. +Qed. + +Remark alloc_globals_nextblock: + forall gl m m', + alloc_globals m gl = Some m' -> + Mem.nextblock m' = advance_next gl (Mem.nextblock m). +Proof. + induction gl; simpl; intros. + congruence. + destruct (alloc_global m a) as [m1|] eqn:?; try discriminate. + erewrite IHgl; eauto. erewrite alloc_global_nextblock; eauto. +Qed. + +(** Permissions *) + +Remark store_zeros_perm: + forall k prm b' q m b p n m', + store_zeros m b p n = Some m' -> + (Mem.perm m b' q k prm <-> Mem.perm m' b' q k prm). +Proof. + intros until n. functional induction (store_zeros m b p n); intros. + inv H; tauto. + destruct (IHo _ H); intros. split; eauto with mem. + congruence. +Qed. + +Remark store_init_data_perm: + forall k prm b' q i b m p m', + store_init_data m b p i = Some m' -> + (Mem.perm m b' q k prm <-> Mem.perm m' b' q k prm). +Proof. + intros. + assert (forall chunk v, + Mem.store chunk m b p v = Some m' -> + (Mem.perm m b' q k prm <-> Mem.perm m' b' q k prm)). + intros; split; eauto with mem. + destruct i; simpl in H; eauto. + inv H; tauto. + destruct (find_symbol ge i); try discriminate. eauto. +Qed. + +Remark store_init_data_list_perm: + forall k prm b' q idl b m p m', + store_init_data_list m b p idl = Some m' -> + (Mem.perm m b' q k prm <-> Mem.perm m' b' q k prm). +Proof. + induction idl as [ | i1 idl]; simpl; intros. +- inv H; tauto. +- destruct (store_init_data m b p i1) as [m1|] eqn:S1; try discriminate. + transitivity (Mem.perm m1 b' q k prm). + eapply store_init_data_perm; eauto. + eapply IHidl; eauto. +Qed. + +Remark alloc_global_perm: + forall k prm b' q idg m m', + alloc_global m idg = Some m' -> + Mem.valid_block m b' -> + (Mem.perm m b' q k prm <-> Mem.perm m' b' q k prm). +Proof. + intros. destruct idg as [id [[f|v]|]]; simpl in H. + (* function *) + destruct (Mem.alloc m 0 1) as [m1 b] eqn:?. + assert (b' <> b). apply Mem.valid_not_valid_diff with m; eauto with mem. + split; intros. + eapply Mem.perm_drop_3; eauto. eapply Mem.perm_alloc_1; eauto. + eapply Mem.perm_alloc_4; eauto. eapply Mem.perm_drop_4; eauto. + (* variable *) + set (init := gvar_init v) in *. + set (sz := init_data_list_size init) in *. + destruct (Mem.alloc m 0 sz) as [m1 b] eqn:?. + destruct (store_zeros m1 b 0 sz) as [m2|] eqn:?; try discriminate. + destruct (store_init_data_list m2 b 0 init) as [m3|] eqn:?; try discriminate. + assert (b' <> b). apply Mem.valid_not_valid_diff with m; eauto with mem. + split; intros. + eapply Mem.perm_drop_3; eauto. + erewrite <- store_init_data_list_perm; [idtac|eauto]. + erewrite <- store_zeros_perm; [idtac|eauto]. + eapply Mem.perm_alloc_1; eauto. + eapply Mem.perm_alloc_4; eauto. + erewrite store_zeros_perm; [idtac|eauto]. + erewrite store_init_data_list_perm; [idtac|eauto]. + eapply Mem.perm_drop_4; eauto. + (* none *) + destruct (Mem.alloc m 0 0) as [m1 b] eqn:?. + assert (b' <> b). apply Mem.valid_not_valid_diff with m; eauto with mem. + inv H. + split; intros. + eapply Mem.perm_alloc_1; eauto. + eapply Mem.perm_alloc_4; eauto. +Qed. + +Remark alloc_globals_perm: + forall k prm b' q gl m m', + alloc_globals m gl = Some m' -> + Mem.valid_block m b' -> + (Mem.perm m b' q k prm <-> Mem.perm m' b' q k prm). +Proof. + induction gl. + simpl; intros. inv H. tauto. + simpl; intros. destruct (alloc_global m a) as [m1|] eqn:?; try discriminate. + erewrite alloc_global_perm; eauto. eapply IHgl; eauto. + unfold Mem.valid_block in *. erewrite alloc_global_nextblock; eauto. + apply Plt_trans_succ; auto. +Qed. + +(** Data preservation properties *) + +Remark store_zeros_unchanged: + forall (P: block -> Z -> Prop) m b p n m', + store_zeros m b p n = Some m' -> + (forall i, p <= i < p + n -> ~ P b i) -> + Mem.unchanged_on P m m'. +Proof. + intros until n. functional induction (store_zeros m b p n); intros. +- inv H; apply Mem.unchanged_on_refl. +- apply Mem.unchanged_on_trans with m'. ++ eapply Mem.store_unchanged_on; eauto. simpl. intros. apply H0. omega. ++ apply IHo; auto. intros; apply H0; omega. +- discriminate. +Qed. + +Remark store_init_data_unchanged: + forall (P: block -> Z -> Prop) b i m p m', + store_init_data m b p i = Some m' -> + (forall ofs, p <= ofs < p + init_data_size i -> ~ P b ofs) -> + Mem.unchanged_on P m m'. +Proof. + intros. destruct i; simpl in *; + try (eapply Mem.store_unchanged_on; eauto; fail). + inv H; apply Mem.unchanged_on_refl. + destruct (find_symbol ge i); try congruence. + eapply Mem.store_unchanged_on; eauto; + unfold Mptr; destruct Archi.ptr64; eauto. +Qed. + +Remark store_init_data_list_unchanged: + forall (P: block -> Z -> Prop) b il m p m', + store_init_data_list m b p il = Some m' -> + (forall ofs, p <= ofs -> ~ P b ofs) -> + Mem.unchanged_on P m m'. +Proof. + induction il; simpl; intros. +- inv H. apply Mem.unchanged_on_refl. +- destruct (store_init_data m b p a) as [m1|] eqn:?; try congruence. + apply Mem.unchanged_on_trans with m1. + eapply store_init_data_unchanged; eauto. intros; apply H0; tauto. + eapply IHil; eauto. intros; apply H0. generalize (init_data_size_pos a); omega. +Qed. + +(** Properties related to [loadbytes] *) + +Definition readbytes_as_zero (m: mem) (b: block) (ofs len: Z) : Prop := + forall p n, + ofs <= p -> p + Z.of_nat n <= ofs + len -> + Mem.loadbytes m b p (Z.of_nat n) = Some (list_repeat n (Byte Byte.zero)). + +Lemma store_zeros_loadbytes: + forall m b p n m', + store_zeros m b p n = Some m' -> + readbytes_as_zero m' b p n. +Proof. + intros until n; functional induction (store_zeros m b p n); red; intros. +- destruct n0. simpl. apply Mem.loadbytes_empty. omega. + rewrite inj_S in H1. omegaContradiction. +- destruct (zeq p0 p). + + subst p0. destruct n0. simpl. apply Mem.loadbytes_empty. omega. + rewrite inj_S in H1. rewrite inj_S. + replace (Z.succ (Z.of_nat n0)) with (1 + Z.of_nat n0) by omega. + change (list_repeat (S n0) (Byte Byte.zero)) + with ((Byte Byte.zero :: nil) ++ list_repeat n0 (Byte Byte.zero)). + apply Mem.loadbytes_concat. + eapply Mem.loadbytes_unchanged_on with (P := fun b1 ofs1 => ofs1 = p). + eapply store_zeros_unchanged; eauto. intros; omega. + intros; omega. + replace (Byte Byte.zero :: nil) with (encode_val Mint8unsigned Vzero). + change 1 with (size_chunk Mint8unsigned). + eapply Mem.loadbytes_store_same; eauto. + unfold encode_val; unfold encode_int; unfold rev_if_be; destruct Archi.big_endian; reflexivity. + eapply IHo; eauto. omega. omega. omega. omega. + + eapply IHo; eauto. omega. omega. +- discriminate. +Qed. + +Definition bytes_of_init_data (i: init_data): list memval := + match i with + | Init_int8 n => inj_bytes (encode_int 1%nat (Int.unsigned n)) + | Init_int16 n => inj_bytes (encode_int 2%nat (Int.unsigned n)) + | Init_int32 n => inj_bytes (encode_int 4%nat (Int.unsigned n)) + | Init_int64 n => inj_bytes (encode_int 8%nat (Int64.unsigned n)) + | Init_float32 n => inj_bytes (encode_int 4%nat (Int.unsigned (Float32.to_bits n))) + | Init_float64 n => inj_bytes (encode_int 8%nat (Int64.unsigned (Float.to_bits n))) + | Init_space n => list_repeat (Z.to_nat n) (Byte Byte.zero) + | Init_addrof id ofs => + match find_symbol ge id with + | Some b => inj_value (if Archi.ptr64 then Q64 else Q32) (Vptr b ofs) + | None => list_repeat (if Archi.ptr64 then 8%nat else 4%nat) Undef + end + end. + +Remark init_data_size_addrof: + forall id ofs, init_data_size (Init_addrof id ofs) = size_chunk Mptr. +Proof. + intros. unfold Mptr. simpl. destruct Archi.ptr64; auto. +Qed. + +Lemma store_init_data_loadbytes: + forall m b p i m', + store_init_data m b p i = Some m' -> + readbytes_as_zero m b p (init_data_size i) -> + Mem.loadbytes m' b p (init_data_size i) = Some (bytes_of_init_data i). +Proof. + intros; destruct i; simpl in H; try apply (Mem.loadbytes_store_same _ _ _ _ _ _ H). +- inv H. simpl. + assert (EQ: Z.of_nat (Z.to_nat z) = Z.max z 0). + { destruct (zle 0 z). rewrite Z2Nat.id; xomega. destruct z; try discriminate. simpl. xomega. } + rewrite <- EQ. apply H0. omega. simpl. omega. +- rewrite init_data_size_addrof. simpl. + destruct (find_symbol ge i) as [b'|]; try discriminate. + rewrite (Mem.loadbytes_store_same _ _ _ _ _ _ H). + unfold encode_val, Mptr; destruct Archi.ptr64; reflexivity. +Qed. + +Fixpoint bytes_of_init_data_list (il: list init_data): list memval := + match il with + | nil => nil + | i :: il => bytes_of_init_data i ++ bytes_of_init_data_list il + end. + +Lemma store_init_data_list_loadbytes: + forall b il m p m', + store_init_data_list m b p il = Some m' -> + readbytes_as_zero m b p (init_data_list_size il) -> + Mem.loadbytes m' b p (init_data_list_size il) = Some (bytes_of_init_data_list il). +Proof. + induction il as [ | i1 il]; simpl; intros. +- apply Mem.loadbytes_empty. omega. +- generalize (init_data_size_pos i1) (init_data_list_size_pos il); intros P1 PL. + destruct (store_init_data m b p i1) as [m1|] eqn:S; try discriminate. + apply Mem.loadbytes_concat. + eapply Mem.loadbytes_unchanged_on with (P := fun b1 ofs1 => ofs1 < p + init_data_size i1). + eapply store_init_data_list_unchanged; eauto. + intros; omega. + intros; omega. + eapply store_init_data_loadbytes; eauto. + red; intros; apply H0. omega. omega. + apply IHil with m1; auto. + red; intros. + eapply Mem.loadbytes_unchanged_on with (P := fun b1 ofs1 => p + init_data_size i1 <= ofs1). + eapply store_init_data_unchanged; eauto. + intros; omega. + intros; omega. + apply H0. omega. omega. + auto. auto. +Qed. + +(** Properties related to [load] *) + +Definition read_as_zero (m: mem) (b: block) (ofs len: Z) : Prop := + forall chunk p, + ofs <= p -> p + size_chunk chunk <= ofs + len -> + (align_chunk chunk | p) -> + Mem.load chunk m b p = + Some (match chunk with + | Mint8unsigned | Mint8signed | Mint16unsigned | Mint16signed | Mint32 => Vint Int.zero + | Mint64 => Vlong Int64.zero + | Mfloat32 => Vsingle Float32.zero + | Mfloat64 => Vfloat Float.zero + | Many32 | Many64 => Vundef + end). + +Remark read_as_zero_unchanged: + forall (P: block -> Z -> Prop) m b ofs len m', + read_as_zero m b ofs len -> + Mem.unchanged_on P m m' -> + (forall i, ofs <= i < ofs + len -> P b i) -> + read_as_zero m' b ofs len. +Proof. + intros; red; intros. eapply Mem.load_unchanged_on; eauto. + intros; apply H1. omega. +Qed. + +Lemma store_zeros_read_as_zero: + forall m b p n m', + store_zeros m b p n = Some m' -> + read_as_zero m' b p n. +Proof. + intros; red; intros. + transitivity (Some(decode_val chunk (list_repeat (size_chunk_nat chunk) (Byte Byte.zero)))). + apply Mem.loadbytes_load; auto. rewrite size_chunk_conv. + eapply store_zeros_loadbytes; eauto. rewrite <- size_chunk_conv; auto. + f_equal. destruct chunk; unfold decode_val; unfold decode_int; unfold rev_if_be; destruct Archi.big_endian; reflexivity. +Qed. + +Fixpoint load_store_init_data (m: mem) (b: block) (p: Z) (il: list init_data) {struct il} : Prop := + match il with + | nil => True + | Init_int8 n :: il' => + Mem.load Mint8unsigned m b p = Some(Vint(Int.zero_ext 8 n)) + /\ load_store_init_data m b (p + 1) il' + | Init_int16 n :: il' => + Mem.load Mint16unsigned m b p = Some(Vint(Int.zero_ext 16 n)) + /\ load_store_init_data m b (p + 2) il' + | Init_int32 n :: il' => + Mem.load Mint32 m b p = Some(Vint n) + /\ load_store_init_data m b (p + 4) il' + | Init_int64 n :: il' => + Mem.load Mint64 m b p = Some(Vlong n) + /\ load_store_init_data m b (p + 8) il' + | Init_float32 n :: il' => + Mem.load Mfloat32 m b p = Some(Vsingle n) + /\ load_store_init_data m b (p + 4) il' + | Init_float64 n :: il' => + Mem.load Mfloat64 m b p = Some(Vfloat n) + /\ load_store_init_data m b (p + 8) il' + | Init_addrof symb ofs :: il' => + (exists b', find_symbol ge symb = Some b' /\ Mem.load Mptr m b p = Some(Vptr b' ofs)) + /\ load_store_init_data m b (p + size_chunk Mptr) il' + | Init_space n :: il' => + read_as_zero m b p n + /\ load_store_init_data m b (p + Zmax n 0) il' + end. + +Lemma store_init_data_list_charact: + forall b il m p m', + store_init_data_list m b p il = Some m' -> + read_as_zero m b p (init_data_list_size il) -> + load_store_init_data m' b p il. +Proof. + assert (A: forall chunk v m b p m1 il m', + Mem.store chunk m b p v = Some m1 -> + store_init_data_list m1 b (p + size_chunk chunk) il = Some m' -> + Mem.load chunk m' b p = Some(Val.load_result chunk v)). + { + intros. + eapply Mem.load_unchanged_on with (P := fun b' ofs' => ofs' < p + size_chunk chunk). + eapply store_init_data_list_unchanged; eauto. intros; omega. + intros; tauto. + eapply Mem.load_store_same; eauto. + } + induction il; simpl. +- auto. +- intros. destruct (store_init_data m b p a) as [m1|] eqn:?; try congruence. + exploit IHil; eauto. + set (P := fun (b': block) ofs' => p + init_data_size a <= ofs'). + apply read_as_zero_unchanged with (m := m) (P := P). + red; intros; apply H0; auto. generalize (init_data_size_pos a); omega. omega. + eapply store_init_data_unchanged with (P := P); eauto. + intros; unfold P. omega. + intros; unfold P. omega. + intro D. + destruct a; simpl in Heqo. ++ split; auto. eapply (A Mint8unsigned (Vint i)); eauto. ++ split; auto. eapply (A Mint16unsigned (Vint i)); eauto. ++ split; auto. eapply (A Mint32 (Vint i)); eauto. ++ split; auto. eapply (A Mint64 (Vlong i)); eauto. ++ split; auto. eapply (A Mfloat32 (Vsingle f)); eauto. ++ split; auto. eapply (A Mfloat64 (Vfloat f)); eauto. ++ split; auto. + set (P := fun (b': block) ofs' => ofs' < p + init_data_size (Init_space z)). + inv Heqo. apply read_as_zero_unchanged with (m := m1) (P := P). + red; intros. apply H0; auto. simpl. generalize (init_data_list_size_pos il); xomega. + eapply store_init_data_list_unchanged; eauto. + intros; unfold P. omega. + intros; unfold P. simpl; xomega. ++ rewrite init_data_size_addrof in *. + split; auto. + destruct (find_symbol ge i); try congruence. + exists b0; split; auto. + transitivity (Some (Val.load_result Mptr (Vptr b0 i0))). + eapply (A Mptr (Vptr b0 i0)); eauto. + unfold Val.load_result, Mptr; destruct Archi.ptr64; auto. +Qed. + +Remark alloc_global_unchanged: + forall (P: block -> Z -> Prop) m id g m', + alloc_global m (id, g) = Some m' -> + Mem.unchanged_on P m m'. +Proof. + intros. destruct g as [[f|v]|]; simpl in H. +- (* function *) + destruct (Mem.alloc m 0 1) as [m1 b] eqn:?. + set (Q := fun b' (ofs: Z) => b' <> b). + apply Mem.unchanged_on_implies with Q. + apply Mem.unchanged_on_trans with m1. + eapply Mem.alloc_unchanged_on; eauto. + eapply Mem.drop_perm_unchanged_on; eauto. + intros; red. apply Mem.valid_not_valid_diff with m; eauto with mem. +- (* variable *) + set (init := gvar_init v) in *. + set (sz := init_data_list_size init) in *. + destruct (Mem.alloc m 0 sz) as [m1 b] eqn:?. + destruct (store_zeros m1 b 0 sz) as [m2|] eqn:?; try discriminate. + destruct (store_init_data_list m2 b 0 init) as [m3|] eqn:?; try discriminate. + set (Q := fun b' (ofs: Z) => b' <> b). + apply Mem.unchanged_on_implies with Q. + apply Mem.unchanged_on_trans with m1. + eapply Mem.alloc_unchanged_on; eauto. + apply Mem.unchanged_on_trans with m2. + eapply store_zeros_unchanged; eauto. + apply Mem.unchanged_on_trans with m3. + eapply store_init_data_list_unchanged; eauto. + eapply Mem.drop_perm_unchanged_on; eauto. + intros; red. apply Mem.valid_not_valid_diff with m; eauto with mem. +- (* none *) + destruct (Mem.alloc m 0 0) as [m1 b] eqn:?. + inv H. + eapply Mem.alloc_unchanged_on; eauto. +Qed. + +Remark alloc_globals_unchanged: + forall (P: block -> Z -> Prop) gl m m', + alloc_globals m gl = Some m' -> + Mem.unchanged_on P m m'. +Proof. + induction gl; simpl; intros. +- inv H. apply Mem.unchanged_on_refl. +- destruct (alloc_global m a) as [m''|] eqn:?; try discriminate. + destruct a as [id g]. + apply Mem.unchanged_on_trans with m''. + eapply alloc_global_unchanged; eauto. + apply IHgl; auto. +Qed. + +Remark load_store_init_data_invariant: + forall m m' b, + (forall chunk ofs, Mem.load chunk m' b ofs = Mem.load chunk m b ofs) -> + forall il p, + load_store_init_data m b p il -> load_store_init_data m' b p il. +Proof. + induction il; intro p; simpl. + auto. + rewrite ! H. destruct a; intuition. red; intros; rewrite H; auto. +Qed. + +Definition globals_initialized (g: t F V) (m: mem) := + forall b gd, + find_def g b = Some gd -> + match gd with + | Gfun f => + Mem.perm m b 0 Cur Nonempty + /\ (forall ofs k p, Mem.perm m b ofs k p -> ofs = 0 /\ p = Nonempty) + | Gvar v => + Mem.range_perm m b 0 (init_data_list_size v.(gvar_init)) Cur (perm_globvar v) + /\ (forall ofs k p, Mem.perm m b ofs k p -> + 0 <= ofs < init_data_list_size v.(gvar_init) /\ perm_order (perm_globvar v) p) + /\ (v.(gvar_volatile) = false -> load_store_init_data m b 0 v.(gvar_init)) + /\ (v.(gvar_volatile) = false -> Mem.loadbytes m b 0 (init_data_list_size v.(gvar_init)) = Some (bytes_of_init_data_list v.(gvar_init))) + end. + +Lemma alloc_global_initialized: + forall g m id gd m', + genv_next g = Mem.nextblock m -> + alloc_global m (id, gd) = Some m' -> + globals_initialized g m -> + globals_initialized (add_global g (id, gd)) m' + /\ genv_next (add_global g (id, gd)) = Mem.nextblock m'. +Proof. + intros. + exploit alloc_global_nextblock; eauto. intros NB. split. +- (* globals-initialized *) + red; intros. unfold find_def in H2; simpl in H2. + destruct gd. +{ + rewrite PTree.gsspec in H2. destruct (peq b (genv_next g)). ++ inv H2. destruct gd0 as [f|v]; simpl in H0. +* destruct (Mem.alloc m 0 1) as [m1 b] eqn:ALLOC. + exploit Mem.alloc_result; eauto. intros RES. + rewrite H, <- RES. split. + eapply Mem.perm_drop_1; eauto. omega. + intros. + assert (0 <= ofs < 1). { eapply Mem.perm_alloc_3; eauto. eapply Mem.perm_drop_4; eauto. } + exploit Mem.perm_drop_2; eauto. intros ORD. + split. omega. inv ORD; auto. +* set (init := gvar_init v) in *. + set (sz := init_data_list_size init) in *. + destruct (Mem.alloc m 0 sz) as [m1 b] eqn:?. + destruct (store_zeros m1 b 0 sz) as [m2|] eqn:?; try discriminate. + destruct (store_init_data_list m2 b 0 init) as [m3|] eqn:?; try discriminate. + exploit Mem.alloc_result; eauto. intro RES. + replace (genv_next g) with b by congruence. + split. red; intros. eapply Mem.perm_drop_1; eauto. + split. intros. + assert (0 <= ofs < sz). + { eapply Mem.perm_alloc_3; eauto. + erewrite store_zeros_perm by eauto. + erewrite store_init_data_list_perm by eauto. + eapply Mem.perm_drop_4; eauto. } + split; auto. + eapply Mem.perm_drop_2; eauto. + split. intros NOTVOL. apply load_store_init_data_invariant with m3. + intros. eapply Mem.load_drop; eauto. right; right; right. + unfold perm_globvar. rewrite NOTVOL. destruct (gvar_readonly v); auto with mem. + eapply store_init_data_list_charact; eauto. + eapply store_zeros_read_as_zero; eauto. + intros NOTVOL. + transitivity (Mem.loadbytes m3 b 0 sz). + eapply Mem.loadbytes_drop; eauto. right; right; right. + unfold perm_globvar. rewrite NOTVOL. destruct (gvar_readonly v); auto with mem. + eapply store_init_data_list_loadbytes; eauto. + eapply store_zeros_loadbytes; eauto. ++ assert (U: Mem.unchanged_on (fun _ _ => True) m m') by (eapply alloc_global_unchanged; eauto). + assert (VALID: Mem.valid_block m b). + { red. rewrite <- H. eapply genv_defs_range; eauto. } + exploit H1; eauto. + destruct gd0 as [f|v]. +* intros [A B]; split; intros. + eapply Mem.perm_unchanged_on; eauto. exact I. + eapply B. eapply Mem.perm_unchanged_on_2; eauto. exact I. +* intros (A & B & C & D). split; [| split; [| split]]. + red; intros. eapply Mem.perm_unchanged_on; eauto. exact I. + intros. eapply B. eapply Mem.perm_unchanged_on_2; eauto. exact I. + intros. apply load_store_init_data_invariant with m; auto. + intros. eapply Mem.load_unchanged_on_1; eauto. intros; exact I. + intros. eapply Mem.loadbytes_unchanged_on; eauto. intros; exact I. +} +{ + generalize H0. intro H0' . + apply alloc_global_nextblock in H0'. + apply (alloc_global_unchanged (fun _ _ => True)) in H0. + generalize H2. intro H2'. + apply genv_defs_range in H2'. + rewrite H in H2'. + apply H1 in H2. + destruct gd0. + + destruct H2. + split. + - eapply Mem.perm_unchanged_on; eauto. + simpl; auto. + - intros. + eapply H3. + eapply Mem.perm_unchanged_on_2; eauto. + simpl; auto. + + destruct H2 as (R & PO & LSINI & INI). + split. + { + red; intros. eapply Mem.perm_unchanged_on; eauto. + simpl; auto. + } + split. + { + intros. + eapply PO; eauto. + eapply Mem.perm_unchanged_on_2; eauto. + simpl; auto. + } + split. + { + intros. + eapply load_store_init_data_invariant. + 2: eapply LSINI; eauto. + intros. + eapply Mem.load_unchanged_on_1; eauto. + simpl; auto. + } + { + intros. + eapply Mem.loadbytes_unchanged_on; eauto. + simpl; auto. + } +} +- simpl. congruence. +Qed. + +Lemma alloc_globals_initialized: + forall gl ge m m', + alloc_globals m gl = Some m' -> + genv_next ge = Mem.nextblock m -> + globals_initialized ge m -> + globals_initialized (add_globals ge gl) m'. +Proof. + induction gl; simpl; intros. +- inv H; auto. +- destruct a as [id g]. destruct (alloc_global m (id, g)) as [m1|] eqn:?; try discriminate. + exploit alloc_global_initialized; eauto. intros [P Q]. + eapply IHgl; eauto. +Qed. + +Definition globals_initialized_strong (g: t F V) (m: mem) := + (forall b, + find_def g b = None -> + forall ofs k p, ~ Mem.perm m b ofs k p) /\ + globals_initialized g m. + +Lemma alloc_global_initialized_strong: + forall g m id gd m', + genv_next g = Mem.nextblock m -> + alloc_global m (id, gd) = Some m' -> + globals_initialized_strong g m -> + globals_initialized_strong (add_global g (id, gd)) m' + /\ genv_next (add_global g (id, gd)) = Mem.nextblock m' . +Proof. + intros g m id gd m' H H0 H1. + exploit alloc_global_nextblock; eauto. intros NB. + exploit (alloc_global_unchanged (fun _ _ => True)); eauto. intro UNCH. + destruct H1 as [H1 H2]. + apply and_assoc. + split. + { + intros b H3 ofs k p. + unfold find_def in H3. + unfold add_global in H3. + simpl in H3. + destruct gd. + + rewrite PTree.gsspec in H3. + destruct (peq b (genv_next g)); try discriminate. + intro PERM. + eapply H1; eauto. + eapply Mem.perm_unchanged_on_2 with (P := fun _ _ => True); simpl; eauto. + unfold Mem.valid_block. + apply Mem.perm_valid_block in PERM. + unfold Mem.valid_block in PERM. + rewrite H in n. + rewrite NB in PERM. + xomega. + + intro PERM. + destruct (peq b (genv_next g)) as [ | n ] . + - subst. + simpl in H0. + destruct (Mem.alloc m 0 0) as [? b] eqn:ALLOC. + inv H0. + exploit Mem.alloc_result; eauto. + intro; subst. + rewrite H in PERM. + exploit Mem.perm_alloc_3; eauto. + omega. + - eapply H1; eauto. + eapply Mem.perm_unchanged_on_2; eauto. + { simpl; auto. } + unfold Mem.valid_block. + rewrite H in n. + apply Mem.perm_valid_block in PERM. + unfold Mem.valid_block in PERM. + rewrite NB in PERM. + xomega. + } + eapply alloc_global_initialized; eauto. +Qed. + +Lemma alloc_globals_initialized_strong: + forall gl ge m m', + alloc_globals m gl = Some m' -> + genv_next ge = Mem.nextblock m -> + globals_initialized_strong ge m -> + globals_initialized_strong (add_globals ge gl) m'. +Proof. + induction gl; simpl; intros. +- inv H; auto. +- destruct a as [id g]. destruct (alloc_global m (id, g)) as [m1|] eqn:?; try discriminate. + exploit alloc_global_initialized_strong; eauto. intros [P Q]. + eapply IHgl; eauto. +Qed. + +End INITMEM. + +Definition init_mem (p: program F V) := + alloc_globals (globalenv p) Mem.empty p.(prog_defs). + +Lemma init_mem_genv_next: forall p m, + init_mem p = Some m -> + genv_next (globalenv p) = Mem.nextblock m. +Proof. + unfold init_mem; intros. + exploit alloc_globals_nextblock; eauto. rewrite Mem.nextblock_empty. intro. + generalize (genv_next_add_globals (prog_defs p) (empty_genv F V (prog_public p))). + fold (globalenv p). simpl genv_next. intros. congruence. +Qed. + +Theorem find_symbol_not_fresh: + forall p id b m, + init_mem p = Some m -> + find_symbol (globalenv p) id = Some b -> Mem.valid_block m b. +Proof. + intros. red. erewrite <- init_mem_genv_next; eauto. + eapply genv_symb_range; eauto. +Qed. + +Theorem find_def_not_fresh: + forall p b g m, + init_mem p = Some m -> + find_def (globalenv p) b = Some g -> Mem.valid_block m b. +Proof. + intros. red. erewrite <- init_mem_genv_next; eauto. + eapply genv_defs_range; eauto. +Qed. + +Theorem find_funct_ptr_not_fresh: + forall p b f m, + init_mem p = Some m -> + find_funct_ptr (globalenv p) b = Some f -> Mem.valid_block m b. +Proof. + intros. rewrite find_funct_ptr_iff in H0. eapply find_def_not_fresh; eauto. +Qed. + +Theorem find_var_info_not_fresh: + forall p b gv m, + init_mem p = Some m -> + find_var_info (globalenv p) b = Some gv -> Mem.valid_block m b. +Proof. + intros. rewrite find_var_info_iff in H0. eapply find_def_not_fresh; eauto. +Qed. + +Lemma init_mem_characterization_gen: + forall p m, + init_mem p = Some m -> + globals_initialized (globalenv p) (globalenv p) m. +Proof. + intros. apply alloc_globals_initialized with Mem.empty. + auto. + rewrite Mem.nextblock_empty. auto. + red; intros. unfold find_def in H0; simpl in H0; rewrite PTree.gempty in H0; discriminate. +Qed. + +Lemma init_mem_characterization_gen_strong: + forall p m, + init_mem p = Some m -> + globals_initialized_strong (globalenv p) (globalenv p) m. +Proof. + intros. apply alloc_globals_initialized_strong with Mem.empty. + auto. + rewrite Mem.nextblock_empty. auto. + red; intros. + split. + { + intros; eauto using Mem.perm_empty. + } + red; intros. + unfold find_def in H0; simpl in H0; rewrite PTree.gempty in H0; discriminate. +Qed. + +Theorem init_mem_characterization: + forall p b gv m, + find_var_info (globalenv p) b = Some gv -> + init_mem p = Some m -> + Mem.range_perm m b 0 (init_data_list_size gv.(gvar_init)) Cur (perm_globvar gv) + /\ (forall ofs k p, Mem.perm m b ofs k p -> + 0 <= ofs < init_data_list_size gv.(gvar_init) /\ perm_order (perm_globvar gv) p) + /\ (gv.(gvar_volatile) = false -> + load_store_init_data (globalenv p) m b 0 gv.(gvar_init)) + /\ (gv.(gvar_volatile) = false -> + Mem.loadbytes m b 0 (init_data_list_size gv.(gvar_init)) = Some (bytes_of_init_data_list (globalenv p) gv.(gvar_init))). +Proof. + intros. rewrite find_var_info_iff in H. + exploit init_mem_characterization_gen; eauto. +Qed. + +Theorem init_mem_characterization_2: + forall p b fd m, + find_funct_ptr (globalenv p) b = Some fd -> + init_mem p = Some m -> + Mem.perm m b 0 Cur Nonempty + /\ (forall ofs k p, Mem.perm m b ofs k p -> ofs = 0 /\ p = Nonempty). +Proof. + intros. rewrite find_funct_ptr_iff in H. + exploit init_mem_characterization_gen; eauto. +Qed. + +(** ** Compatibility with memory injections *) + +Section INITMEM_INJ. + +Variable ge: t F V. +Variable thr: block. +Hypothesis symb_inject: forall id b, find_symbol ge id = Some b -> Plt b thr. + +Lemma store_zeros_neutral: + forall m b p n m', + Mem.inject_neutral thr m -> + Plt b thr -> + store_zeros m b p n = Some m' -> + Mem.inject_neutral thr m'. +Proof. + intros until n. functional induction (store_zeros m b p n); intros. + inv H1; auto. + apply IHo; auto. eapply Mem.store_inject_neutral; eauto. constructor. + inv H1. +Qed. + +Lemma store_init_data_neutral: + forall m b p id m', + Mem.inject_neutral thr m -> + Plt b thr -> + store_init_data ge m b p id = Some m' -> + Mem.inject_neutral thr m'. +Proof. + intros. + destruct id; simpl in H1; try (eapply Mem.store_inject_neutral; eauto; fail). + congruence. + destruct (find_symbol ge i) as [b'|] eqn:E; try discriminate. + eapply Mem.store_inject_neutral; eauto. + econstructor. unfold Mem.flat_inj. apply pred_dec_true; auto. eauto. + rewrite Ptrofs.add_zero. auto. +Qed. + +Lemma store_init_data_list_neutral: + forall b idl m p m', + Mem.inject_neutral thr m -> + Plt b thr -> + store_init_data_list ge m b p idl = Some m' -> + Mem.inject_neutral thr m'. +Proof. + induction idl; simpl; intros. + congruence. + destruct (store_init_data ge m b p a) as [m1|] eqn:E; try discriminate. + eapply IHidl. eapply store_init_data_neutral; eauto. auto. eauto. +Qed. + +Lemma alloc_global_neutral: + forall idg m m', + alloc_global ge m idg = Some m' -> + Mem.inject_neutral thr m -> + Plt (Mem.nextblock m) thr -> + Mem.inject_neutral thr m'. +Proof. + intros. destruct idg as [id [[f|v]|]]; simpl in H. + (* function *) + destruct (Mem.alloc m 0 1) as [m1 b] eqn:?. + assert (Plt b thr). rewrite (Mem.alloc_result _ _ _ _ _ Heqp). auto. + eapply Mem.drop_inject_neutral; eauto. + eapply Mem.alloc_inject_neutral; eauto. + (* variable *) + set (init := gvar_init v) in *. + set (sz := init_data_list_size init) in *. + destruct (Mem.alloc m 0 sz) as [m1 b] eqn:?. + destruct (store_zeros m1 b 0 sz) as [m2|] eqn:?; try discriminate. + destruct (store_init_data_list ge m2 b 0 init) as [m3|] eqn:?; try discriminate. + assert (Plt b thr). rewrite (Mem.alloc_result _ _ _ _ _ Heqp). auto. + eapply Mem.drop_inject_neutral; eauto. + eapply store_init_data_list_neutral with (m := m2) (b := b); eauto. + eapply store_zeros_neutral with (m := m1); eauto. + eapply Mem.alloc_inject_neutral; eauto. + (* none *) + destruct (Mem.alloc m 0 0) as [m1 b] eqn:?. + assert (Plt b thr). rewrite (Mem.alloc_result _ _ _ _ _ Heqp). auto. + inv H. + eapply Mem.alloc_inject_neutral; eauto. +Qed. + +Remark advance_next_le: forall gl x, Ple x (advance_next (F:=F) (V:=V) gl x). +Proof. + induction gl; simpl; intros. + apply Ple_refl. + apply Ple_trans with (Psucc x). apply Ple_succ. eauto. +Qed. + +Lemma alloc_globals_neutral: + forall gl m m', + alloc_globals ge m gl = Some m' -> + Mem.inject_neutral thr m -> + Ple (Mem.nextblock m') thr -> + Mem.inject_neutral thr m'. +Proof. + induction gl; intros. + simpl in *. congruence. + exploit alloc_globals_nextblock; eauto. intros EQ. + simpl in *. destruct (alloc_global ge m a) as [m1|] eqn:E; try discriminate. + exploit alloc_global_neutral; eauto. + assert (Ple (Psucc (Mem.nextblock m)) (Mem.nextblock m')). + { rewrite EQ. apply advance_next_le. } + unfold Plt, Ple in *; zify; omega. +Qed. + +End INITMEM_INJ. + +Theorem initmem_inject: + forall p m, + init_mem p = Some m -> + Mem.inject (Mem.flat_inj (Mem.nextblock m)) m m. +Proof. + unfold init_mem; intros. + apply Mem.neutral_inject. + eapply alloc_globals_neutral; eauto. + intros. exploit find_symbol_not_fresh; eauto. + apply Mem.empty_inject_neutral. + apply Ple_refl. +Qed. + +(** ** Sufficient and necessary conditions for the initial memory to exist. *) + +(** Alignment properties *) + +Definition init_data_alignment (i: init_data) : Z := + match i with + | Init_int8 n => 1 + | Init_int16 n => 2 + | Init_int32 n => 4 + | Init_int64 n => 8 + | Init_float32 n => 4 + | Init_float64 n => 4 + | Init_addrof symb ofs => if Archi.ptr64 then 8 else 4 + | Init_space n => 1 + end. + +Fixpoint init_data_list_aligned (p: Z) (il: list init_data) {struct il} : Prop := + match il with + | nil => True + | i1 :: il => (init_data_alignment i1 | p) /\ init_data_list_aligned (p + init_data_size i1) il + end. + +Section INITMEM_INVERSION. + +Variable ge: t F V. + +Lemma store_init_data_aligned: + forall m b p i m', + store_init_data ge m b p i = Some m' -> + (init_data_alignment i | p). +Proof. + intros. + assert (DFL: forall chunk v, + Mem.store chunk m b p v = Some m' -> + align_chunk chunk = init_data_alignment i -> + (init_data_alignment i | p)). + { intros. apply Mem.store_valid_access_3 in H0. destruct H0. congruence. } + destruct i; simpl in H; eauto. + simpl. apply Z.divide_1_l. + destruct (find_symbol ge i); try discriminate. eapply DFL. eassumption. + unfold Mptr, init_data_alignment; destruct Archi.ptr64; auto. +Qed. + +Lemma store_init_data_list_aligned: + forall b il m p m', + store_init_data_list ge m b p il = Some m' -> + init_data_list_aligned p il. +Proof. + induction il as [ | i1 il]; simpl; intros. +- auto. +- destruct (store_init_data ge m b p i1) as [m1|] eqn:S1; try discriminate. + split; eauto. eapply store_init_data_aligned; eauto. +Qed. + +Lemma store_init_data_list_free_idents: + forall b i o il m p m', + store_init_data_list ge m b p il = Some m' -> + In (Init_addrof i o) il -> + exists b', find_symbol ge i = Some b'. +Proof. + induction il as [ | i1 il]; simpl; intros. +- contradiction. +- destruct (store_init_data ge m b p i1) as [m1|] eqn:S1; try discriminate. + destruct H0. ++ subst i1. simpl in S1. destruct (find_symbol ge i) as [b'|]. exists b'; auto. discriminate. ++ eapply IHil; eauto. +Qed. + +End INITMEM_INVERSION. + +Theorem init_mem_inversion: + forall p m id v, + init_mem p = Some m -> + In (id, Some (Gvar v)) p.(prog_defs) -> + init_data_list_aligned 0 v.(gvar_init) + /\ forall i o, In (Init_addrof i o) v.(gvar_init) -> exists b, find_symbol (globalenv p) i = Some b. +Proof. + intros until v. unfold init_mem. set (ge := globalenv p). + revert m. generalize Mem.empty. generalize (prog_defs p). + induction l as [ | idg1 defs ]; simpl; intros m m'; intros. +- contradiction. +- destruct (alloc_global ge m idg1) as [m''|] eqn:A; try discriminate. + destruct H0. ++ subst idg1; simpl in A. + set (il := gvar_init v) in *. set (sz := init_data_list_size il) in *. + destruct (Mem.alloc m 0 sz) as [m1 b]. + destruct (store_zeros m1 b 0 sz) as [m2|]; try discriminate. + destruct (store_init_data_list ge m2 b 0 il) as [m3|] eqn:B; try discriminate. + split. eapply store_init_data_list_aligned; eauto. intros; eapply store_init_data_list_free_idents; eauto. ++ eapply IHdefs; eauto. +Qed. + +Section INITMEM_EXISTS. + +Variable ge: t F V. + +Lemma store_zeros_exists: + forall m b p n, + Mem.range_perm m b p (p + n) Cur Writable -> + exists m', store_zeros m b p n = Some m'. +Proof. + intros until n. functional induction (store_zeros m b p n); intros PERM. +- exists m; auto. +- apply IHo. red; intros. eapply Mem.perm_store_1; eauto. apply PERM. omega. +- destruct (Mem.valid_access_store m Mint8unsigned b p Vzero) as (m' & STORE). + split. red; intros. apply Mem.perm_cur. apply PERM. simpl in H. omega. + simpl. apply Z.divide_1_l. + congruence. +Qed. + +Lemma store_init_data_exists: + forall m b p i, + Mem.range_perm m b p (p + init_data_size i) Cur Writable -> + (init_data_alignment i | p) -> + (forall id ofs, i = Init_addrof id ofs -> exists b, find_symbol ge id = Some b) -> + exists m', store_init_data ge m b p i = Some m'. +Proof. + intros. + assert (DFL: forall chunk v, + init_data_size i = size_chunk chunk -> + init_data_alignment i = align_chunk chunk -> + exists m', Mem.store chunk m b p v = Some m'). + { intros. destruct (Mem.valid_access_store m chunk b p v) as (m' & STORE). + split. rewrite <- H2; auto. rewrite <- H3; auto. + exists m'; auto. } + destruct i; eauto. + simpl. exists m; auto. + simpl. exploit H1; eauto. intros (b1 & FS). rewrite FS. eapply DFL. + unfold init_data_size, Mptr. destruct Archi.ptr64; auto. + unfold init_data_alignment, Mptr. destruct Archi.ptr64; auto. +Qed. + +Lemma store_init_data_list_exists: + forall b il m p, + Mem.range_perm m b p (p + init_data_list_size il) Cur Writable -> + init_data_list_aligned p il -> + (forall id ofs, In (Init_addrof id ofs) il -> exists b, find_symbol ge id = Some b) -> + exists m', store_init_data_list ge m b p il = Some m'. +Proof. + induction il as [ | i1 il ]; simpl; intros. +- exists m; auto. +- destruct H0. + destruct (@store_init_data_exists m b p i1) as (m1 & S1); eauto. + red; intros. apply H. generalize (init_data_list_size_pos il); omega. + rewrite S1. + apply IHil; eauto. + red; intros. erewrite <- store_init_data_perm by eauto. apply H. generalize (init_data_size_pos i1); omega. +Qed. + +Lemma alloc_global_exists: + forall m idg, + match idg with + | (id, None) => True + | (id, Some (Gfun f)) => True + | (id, Some (Gvar v)) => + init_data_list_aligned 0 v.(gvar_init) + /\ forall i o, In (Init_addrof i o) v.(gvar_init) -> exists b, find_symbol ge i = Some b + end -> + exists m', alloc_global ge m idg = Some m'. +Proof. + intros m [id [[f|v]|]]; intros; simpl. +- destruct (Mem.alloc m 0 1) as [m1 b] eqn:ALLOC. + destruct (Mem.range_perm_drop_2 m1 b 0 1 Nonempty) as [m2 DROP]. + red; intros; eapply Mem.perm_alloc_2; eauto. + exists m2; auto. +- destruct H as [P Q]. + set (sz := init_data_list_size (gvar_init v)). + destruct (Mem.alloc m 0 sz) as [m1 b] eqn:ALLOC. + assert (P1: Mem.range_perm m1 b 0 sz Cur Freeable) by (red; intros; eapply Mem.perm_alloc_2; eauto). + destruct (@store_zeros_exists m1 b 0 sz) as [m2 ZEROS]. + red; intros. apply Mem.perm_implies with Freeable; auto with mem. + rewrite ZEROS. + assert (P2: Mem.range_perm m2 b 0 sz Cur Freeable). + { red; intros. erewrite <- store_zeros_perm by eauto. eauto. } + destruct (@store_init_data_list_exists b (gvar_init v) m2 0) as [m3 STORE]; auto. + red; intros. apply Mem.perm_implies with Freeable; auto with mem. + rewrite STORE. + assert (P3: Mem.range_perm m3 b 0 sz Cur Freeable). + { red; intros. erewrite <- store_init_data_list_perm by eauto. eauto. } + destruct (Mem.range_perm_drop_2 m3 b 0 sz (perm_globvar v)) as [m4 DROP]; auto. + exists m4; auto. +- destruct (Mem.alloc _ _ _); eauto. +Qed. + +End INITMEM_EXISTS. + +Theorem init_mem_exists: + forall p, + (forall id v, In (id, Some (Gvar v)) (prog_defs p) -> + init_data_list_aligned 0 v.(gvar_init) + /\ forall i o, In (Init_addrof i o) v.(gvar_init) -> exists b, find_symbol (globalenv p) i = Some b) -> + exists m, init_mem p = Some m. +Proof. + intros. set (ge := globalenv p) in *. + unfold init_mem. revert H. generalize (prog_defs p) Mem.empty. + induction l as [ | idg l]; simpl; intros. +- exists m; auto. +- destruct (@alloc_global_exists ge m idg) as [m1 A1]. + destruct idg as [id [[f|v]|]]; eauto. + fold ge. rewrite A1. eapply IHl; eauto. +Qed. + +End WITHMEMORYMODEL. + +(** * Commutation with program transformations *) + +Section MATCH_GENVS. + +Context {A B V W: Type} (R: globdef A V -> globdef B W -> Prop). + +Record match_genvs (ge1: t A V) (ge2: t B W): Prop := { + mge_next: + genv_next ge2 = genv_next ge1; + mge_symb: + forall id, PTree.get id (genv_symb ge2) = PTree.get id (genv_symb ge1); + mge_defs: + forall b, option_rel R (PTree.get b (genv_defs ge1)) (PTree.get b (genv_defs ge2)) +}. + +Lemma add_global_match: + forall ge1 ge2 id g1 g2, + match_genvs ge1 ge2 -> + option_rel R g1 g2 -> + match_genvs (add_global ge1 (id, g1)) (add_global ge2 (id, g2)). +Proof. + intros. destruct H. constructor; simpl; intros. +- congruence. +- rewrite mge_next0, ! PTree.gsspec. destruct (peq id0 id); auto. +- + inv H0; auto. + rewrite mge_next0, ! PTree.gsspec. destruct (peq b (genv_next ge1)). + constructor; auto. + auto. +Qed. + +Lemma add_globals_match: + forall gl1 gl2, + list_forall2 (fun idg1 idg2 => fst idg1 = fst idg2 /\ option_rel R (snd idg1) (snd idg2)) gl1 gl2 -> + forall ge1 ge2, match_genvs ge1 ge2 -> + match_genvs (add_globals ge1 gl1) (add_globals ge2 gl2). +Proof. + induction 1; intros; simpl. + auto. + destruct a1 as [id1 g1]; destruct b1 as [id2 g2]; simpl in *; destruct H; subst id2. + apply IHlist_forall2. apply add_global_match; auto. +Qed. + +End MATCH_GENVS. + +Section MATCH_PROGRAMS. + +Context {C F1 V1 F2 V2: Type} {LC: Linker C} {LF: Linker F1} {LV: Linker V1}. +Variable match_fundef: C -> F1 -> F2 -> Prop. +Variable match_varinfo: V1 -> V2 -> Prop. +Variable ctx: C. +Variable p: program F1 V1. +Variable tp: program F2 V2. +Hypothesis progmatch: match_program_gen match_fundef match_varinfo ctx p tp. + +Lemma globalenvs_match: + match_genvs (match_globdef match_fundef match_varinfo ctx) (globalenv p) (globalenv tp). +Proof. + intros. apply add_globals_match. apply progmatch. + constructor; simpl; intros; auto. rewrite ! PTree.gempty. constructor. +Qed. + +Theorem find_def_match_2: + forall b, option_rel (match_globdef match_fundef match_varinfo ctx) + (find_def (globalenv p) b) (find_def (globalenv tp) b). +Proof (mge_defs globalenvs_match). + +Theorem find_def_match: + forall b g, + find_def (globalenv p) b = Some g -> + exists tg, + find_def (globalenv tp) b = Some tg /\ match_globdef match_fundef match_varinfo ctx g tg. +Proof. + intros. generalize (find_def_match_2 b). rewrite H; intros R; inv R. + exists y; auto. +Qed. + +Theorem find_funct_ptr_match: + forall b f, + find_funct_ptr (globalenv p) b = Some f -> + exists cunit tf, + find_funct_ptr (globalenv tp) b = Some tf /\ match_fundef cunit f tf /\ linkorder cunit ctx. +Proof. + intros. rewrite find_funct_ptr_iff in *. apply find_def_match in H. + destruct H as (tg & P & Q). inv Q. + exists ctx', f2; intuition auto. apply find_funct_ptr_iff; auto. +Qed. + +Theorem find_funct_match: + forall v f, + find_funct (globalenv p) v = Some f -> + exists cunit tf, + find_funct (globalenv tp) v = Some tf /\ match_fundef cunit f tf /\ linkorder cunit ctx. +Proof. + intros. exploit find_funct_inv; eauto. intros [b EQ]. subst v. + rewrite find_funct_find_funct_ptr in H. + rewrite find_funct_find_funct_ptr. + apply find_funct_ptr_match. auto. +Qed. + +Theorem find_var_info_match: + forall b v, + find_var_info (globalenv p) b = Some v -> + exists tv, + find_var_info (globalenv tp) b = Some tv /\ match_globvar match_varinfo v tv. +Proof. + intros. rewrite find_var_info_iff in *. apply find_def_match in H. + destruct H as (tg & P & Q). inv Q. + exists v2; split; auto. apply find_var_info_iff; auto. +Qed. + +Theorem find_symbol_match: + forall (s : ident), + find_symbol (globalenv tp) s = find_symbol (globalenv p) s. +Proof. + intros. destruct globalenvs_match. apply mge_symb0. +Qed. + +Theorem genv_next_match: + genv_next (globalenv tp) = genv_next (globalenv p). +Proof. + intros. destruct globalenvs_match. apply mge_next0. +Qed. + +Theorem senv_match: + Senv.equiv (to_senv (globalenv p)) (to_senv (globalenv tp)). +Proof. + red; simpl. repeat split. +- apply genv_next_match. +- apply find_symbol_match. +- intros. unfold public_symbol. rewrite find_symbol_match. + rewrite ! globalenv_public. + destruct progmatch as (P & Q & R). rewrite R. auto. +- intros. unfold block_is_volatile. + destruct globalenvs_match as [P Q R]. specialize (R b). + unfold find_var_info, find_def. + inv R; auto. + inv H1; auto. + inv H2; auto. +Qed. + +Context `{memory_model_prf: Mem.MemoryModel}. + +Lemma store_init_data_list_match: + forall idl m b ofs m', + store_init_data_list (globalenv p) m b ofs idl = Some m' -> + store_init_data_list (globalenv tp) m b ofs idl = Some m'. +Proof. + induction idl; simpl; intros. +- auto. +- destruct (store_init_data (globalenv p) m b ofs a) as [m1|] eqn:S; try discriminate. + assert (X: store_init_data (globalenv tp) m b ofs a = Some m1). + { destruct a; auto. simpl; rewrite find_symbol_match; auto. } + rewrite X. auto. +Qed. + +Lemma alloc_globals_match: + forall gl1 gl2, list_forall2 (match_ident_option_globdef match_fundef match_varinfo ctx) gl1 gl2 -> + forall m m', + alloc_globals (globalenv p) m gl1 = Some m' -> + alloc_globals (globalenv tp) m gl2 = Some m'. +Proof. + induction 1; simpl; intros. +- auto. +- destruct (alloc_global (globalenv p) m a1) as [m1|] eqn:?; try discriminate. + assert (X: alloc_global (globalenv tp) m b1 = Some m1). + { destruct a1 as [id1 g1]; destruct b1 as [id2 g2]; destruct H; simpl in *. + subst id2. inv H2. + - auto. + - inv H; simpl in *. + { auto. } + inv H2; simpl in *. + set (sz := init_data_list_size init) in *. + destruct (Mem.alloc m 0 sz) as [m2 b] eqn:?. + destruct (store_zeros m2 b 0 sz) as [m3|] eqn:?; try discriminate. + destruct (store_init_data_list (globalenv p) m3 b 0 init) as [m4|] eqn:?; try discriminate. + erewrite store_init_data_list_match; eauto. + } + rewrite X; eauto. +Qed. + +Theorem init_mem_match: + forall m, init_mem p = Some m -> init_mem tp = Some m. +Proof. + unfold init_mem; intros. + eapply alloc_globals_match; eauto. apply progmatch. +Qed. + +End MATCH_PROGRAMS. + +(** Special case for partial transformations that do not depend on the compilation unit *) + +Section TRANSFORM_PARTIAL. + +Context {A B V: Type} {LA: Linker A} {LV: Linker V}. +Context {transf: A -> res B} {p: program A V} {tp: program B V}. +Hypothesis progmatch: match_program (fun cu f tf => transf f = OK tf) eq p tp. + +Theorem find_funct_ptr_transf_partial: + forall b f, + find_funct_ptr (globalenv p) b = Some f -> + exists tf, + find_funct_ptr (globalenv tp) b = Some tf /\ transf f = OK tf. +Proof. + intros. exploit (find_funct_ptr_match progmatch); eauto. + intros (cu & tf & P & Q & R); exists tf; auto. +Qed. + +Theorem find_funct_transf_partial: + forall v f, + find_funct (globalenv p) v = Some f -> + exists tf, + find_funct (globalenv tp) v = Some tf /\ transf f = OK tf. +Proof. + intros. exploit (find_funct_match progmatch); eauto. + intros (cu & tf & P & Q & R); exists tf; auto. +Qed. + +Theorem find_symbol_transf_partial: + forall (s : ident), + find_symbol (globalenv tp) s = find_symbol (globalenv p) s. +Proof. + intros. eapply (find_symbol_match progmatch). +Qed. + +Theorem senv_transf_partial: + Senv.equiv (to_senv (globalenv p)) (to_senv (globalenv tp)). +Proof. + intros. eapply (senv_match progmatch). +Qed. + +Context `{memory_model_prf: Mem.MemoryModel}. + +Theorem init_mem_transf_partial: + forall m, init_mem p = Some m -> init_mem tp = Some m. +Proof. + eapply (init_mem_match progmatch). +Qed. + +End TRANSFORM_PARTIAL. + +(** Special case for total transformations that do not depend on the compilation unit *) + +Section TRANSFORM_TOTAL. + +Context {A B V: Type} {LA: Linker A} {LV: Linker V}. +Context {transf: A -> B} {p: program A V} {tp: program B V}. +Hypothesis progmatch: match_program (fun cu f tf => tf = transf f) eq p tp. + +Theorem find_funct_ptr_transf: + forall b f, + find_funct_ptr (globalenv p) b = Some f -> + find_funct_ptr (globalenv tp) b = Some (transf f). +Proof. + intros. exploit (find_funct_ptr_match progmatch); eauto. + intros (cu & tf & P & Q & R). congruence. +Qed. + +Theorem find_funct_transf: + forall v f, + find_funct (globalenv p) v = Some f -> + find_funct (globalenv tp) v = Some (transf f). +Proof. + intros. exploit (find_funct_match progmatch); eauto. + intros (cu & tf & P & Q & R). congruence. +Qed. + +Theorem find_symbol_transf: + forall (s : ident), + find_symbol (globalenv tp) s = find_symbol (globalenv p) s. +Proof. + intros. eapply (find_symbol_match progmatch). +Qed. + +Theorem senv_transf: + Senv.equiv (to_senv (globalenv p)) (to_senv (globalenv tp)). +Proof. + intros. eapply (senv_match progmatch). +Qed. + +Context `{memory_model_prf: Mem.MemoryModel}. + +Theorem init_mem_transf: + forall m, init_mem p = Some m -> init_mem tp = Some m. +Proof. + eapply (init_mem_match progmatch). +Qed. + +End TRANSFORM_TOTAL. +*) +End Genv. + +(* +Coercion Genv.to_senv: Genv.t >-> Senv.t. +*) + diff --git a/src/backend/Environments/HashEnv.v b/src/backend/Environments/HashEnv.v new file mode 100644 index 0000000..12d9a8a --- /dev/null +++ b/src/backend/Environments/HashEnv.v @@ -0,0 +1,62 @@ +Require Import cclib.Coqlib. +Require Import backend.Ctypes. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.Values.LowValues. +Require Import backend.AST. +Require Import backend.IndexLib. + +Inductive HashKey := +| singleton : int256 -> HashKey +| pair : HashKey -> int256 -> HashKey. + +Module HashKeyIndexed <: INDEXED_TYPE. + + Definition t := HashKey. + + Fixpoint index (hk : HashKey) : positive := + match hk with + | singleton i => int_index i + | pair hk i => inject_positive (index hk) (int_index i) + end. + +Definition index_inv (p : positive) : HashKey. +Admitted. + +Lemma index_invertible : forall x, index_inv (index x) = x. +Admitted. + +Lemma index_inj: forall (x y: HashKey), index x = index y -> x = y. +Proof. +Admitted. + +Lemma eq: forall (x y: HashKey), {x = y} + {x <> y}. +Proof. + induction x; destruct y; intros; decide equality; + try decide equality; try apply Int256.eq_dec. +Qed. + +End HashKeyIndexed. + +Module HashKeyMap := IMap(HashKeyIndexed). + +Definition hash_env : Type := HashKeyMap.t val. + +Notation "he [ i ]" := (HashKeyMap.get i he) (at level 80). +Notation "he [ i |-> v ]" := (HashKeyMap.set i v he) (at level 80). + +Definition empty_hash_env : hash_env := + (HashKeyMap.init Vzero). + +Fixpoint hkOfVal (v:val) : option HashKey := + match v with + | Vhash (Vint i) => Some (singleton i) + | Vhash2 v (Vint i) => option_map (fun hk => pair hk i) (hkOfVal v) + | _ => None + end. + +Definition read (ptr:val) (he:hash_env) : option val := + option_map (fun hk => he[hk]) (hkOfVal ptr). + +Definition write (ptr:val) (v:val) (he:hash_env) : option hash_env := + option_map (fun hk => he[hk |-> v]) (hkOfVal ptr). diff --git a/src/backend/Environments/StackEnv.v b/src/backend/Environments/StackEnv.v new file mode 100644 index 0000000..6ccd10b --- /dev/null +++ b/src/backend/Environments/StackEnv.v @@ -0,0 +1,114 @@ +Require Import cclib.Coqlib. +Require Import backend.Ctypes. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.Values.LowValues. +Require Import backend.AST. + +Section StackEnv. + +Variable function : Type. +Variable fn_locals : function -> list (ident * type). + +(* at a lower level, storage and memory are separate *) +Definition stack_env : Type := Int256Map.t LowValues.val. + +(* Memory layout: + 0x00 - 0x1f Return Values + 0x20 - 0x3f Stack Pointer + 0x40 - Stack *) + +Definition sp := Int256.repr (2304 + 32). +Definition sb := Int256.repr (2304 + 64). + +Notation "stk [ i ]" := (Int256Map.get i stk) (at level 80). +Notation "stk [ i |-> v ]" := (Int256Map.set i v stk) (at level 80). + +(* The stack itself begins at 0x40 *) +Definition empty_stack : stack_env := + (Int256Map.init Vzero)[sp|->Vint sb]. + +Definition push_frame (size:int256) (stk:stack_env) : option stack_env := + match stk[sp] with + | Vint p => Some (stk[sp |-> Vint (Int256.add p size)]) + | _ => None + end. + +Definition pop_frame (size:int256) (stk:stack_env) : option stack_env := + match stk[sp] with + | Vint p => Some (stk[sp |-> Vint (Int256.sub p size)]) + | _ => None + end. + +Fixpoint sizeof (ty:type) : int256 := + Int256.repr (sizeof_words ty). + +Fixpoint offset (fl:fieldlist) (id:ident) : option int256 := + match struct_field fl id with + | Some (off,_) => Some (Int256.mul (Int256.repr 32) (Int256.repr (Z.of_nat off))) + | None => None + end. + +Fixpoint frame_size (locs:list (ident * type)) : int256 := + match locs with + | nil => Int256.repr 0 + | cons (_,ty) locs' => Int256.add (Int256.mul (Int256.repr 32) (sizeof ty)) (frame_size locs') + end. + +Definition push_func (f:function) (stk:stack_env) : option stack_env := + push_frame (frame_size (fn_locals f)) stk. + +Definition pop_func (f:function) (stk:stack_env) : option stack_env := + pop_frame (frame_size (fn_locals f)) stk. + +Definition read (ptr:val) (stk:stack_env) : option val := + match ptr,stk[sp] with + | Vint p, Vint s => + if Int256.eq p sp then None + else if Int256.lt p s then Some (stk[p]) + else None + | _,_ => None + end. + +Definition write (ptr:val) (v:val) (stk:stack_env) : option stack_env := + match ptr,stk[sp] with + | Vint p, Vint s => + if Int256.eq p sp then None + else if Int256.lt p s then Some (stk[p |-> v]) + else None + | _,_ => None + end. + + +Fixpoint mkfieldlist (itl:list (ident*type)) : fieldlist := + match itl with + | nil => Fnil + | cons (i,t) itl' => Fcons i t (mkfieldlist itl') + end. + +Definition var (id:ident) (locals:list(ident*type)) (stk:stack_env) : option val := + match stk[sp] with + | Vint s => + let base := Int256.sub s (frame_size locals) in + let fl := mkfieldlist locals in + option_map (fun off => Vint (Int256.add base off)) (offset fl id) + | _ => None + end. + +Definition access_field (v:val) (ty:type) (i:ident) : option val := + match v,ty with + | Vint base, Tstruct _ fl => option_map (fun off => (Vint (Int256.add base off))) (offset fl i) + | _,_ => None + end. + +Definition index_array (v:val) (ty:type) (ind:val) : option val := + match v,ind,ty with + | Vint base, Vint off, Tarray ty' size => + if Int256.lt off (Int256.repr size) then + Some (Vint (Int256.add base (Int256.mul off (sizeof ty')))) + else None + | _, _, _ => None + end. + + +End StackEnv. diff --git a/src/backend/Events.v b/src/backend/Events.v new file mode 100755 index 0000000..b87990b --- /dev/null +++ b/src/backend/Events.v @@ -0,0 +1,17 @@ +(** Observable events, execution traces, and semantics of external calls. *) + +Require Import cclib.Coqlib. +Require Import backend.AST. +Require Import cclib.Integers. +Require Import backend.Values.HighValues. + +(* TODO: make this a list of log entries, method calls *) +Inductive event : Type := + | Etransfer: forall (addr value : val), event + | Ecallmethod: forall (addr : val) (signature: int) (value : val) (args: list val), event + | Elog: forall (topics : list val) (args : list val), event. + +Definition log := list event. + +(* default log. *) +Definition L0 : log := nil. diff --git a/src/backend/Expressions/ExpMiniC.v b/src/backend/Expressions/ExpMiniC.v new file mode 100644 index 0000000..86258c5 --- /dev/null +++ b/src/backend/Expressions/ExpMiniC.v @@ -0,0 +1,99 @@ +(* This is the top level langauge. + + It is almost unchanged from Clight, but there are still quite a few changes. *) + +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** The Clight language: a simplified version of Compcert C where all + expressions are pure and assignments and function calls are + statements, not expressions. *) + +Require Import cclib.Coqlib. +Require Import cclib.Errors. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.AST. +Require Import backend.Ctypes. +Require Import backend.Cop. +Require Import backend.MachineModel. +Require Import backend.Values.HighValues. + +(** * Abstract syntax *) + +(** ** Expressions *) + +(** Clight expressions correspond to the "pure" subset of C expressions. + The main omissions are string literals and assignment operators + ([=], [+=], [++], etc). In Clight, assignment is a statement, + not an expression. Additionally, an expression can also refer to + temporary variables, which are a separate class of local variables + that do not reside in memory and whose address cannot be taken. + + As in Compcert C, all expressions are annotated with their types, + as needed to resolve operator overloading and type-dependent behaviors. *) + + +(* Compared to Clight, we have added array-deref, hash-deref, and calling builtins. *) +Inductive expr : Type := + | Econst_int: int -> type -> expr (** r integer literal *) + | Econst_int256: int256 -> type -> expr (** r 256-bit integer literal *) + | Evar: ident -> type -> expr (** r variable *) + | Eglob: ident -> type -> expr (** global variable *) + | Etempvar: ident -> type -> expr (** r temporary variable *) + | Ederef: expr -> type -> expr (** r pointer dereference (unary [*]) *) + | Eaddr: expr -> type -> expr (** take the address of an lvalue to produce an rvalue pointer *) + | Eunop: unary_operation -> expr -> type -> expr (** r unary operation *) + | Ebinop: binary_operation -> expr -> expr -> type -> expr (** r binary operation *) + | Efield: expr -> ident -> type -> expr (** r access to a member of a struct *) + | Eindex: expr -> expr -> type -> expr + (* Some function calls are compiled to special commands + and have no side effects *) + | Ecall0: builtin0 -> type -> expr + | Ecall1: builtin1 -> expr -> type -> expr. + + +Remark expr_dec: forall (x y: expr), {x = y} + {x <> y}. +Proof. intros. decide equality; +try apply Int256.eq_dec; try apply Int.eq_dec; +try apply type_eq; try apply ident_eq; try decide equality. +Qed. + +(** [sizeof] and [alignof] are derived forms. *) + +(* TODO: if these are ever used by the middle-end, it should probably be changed to Econst_int256. *) +(* +Definition Esizeof (ty' ty: type) : expr := Econst_int (Int.repr(sizeof ty')) ty. +Definition Ealignof (ty' ty: type) : expr := Econst_int (Int.repr(alignof ty')) ty. +*) + +(** Extract the type part of a type-annotated Clight expression. *) + +Definition typeof (e: expr) : type := + match e with + | Econst_int _ ty => ty + | Econst_int256 _ ty => ty + | Evar _ ty => ty + | Eglob _ ty => ty + | Etempvar _ ty => ty + | Ederef _ ty => ty + | Eaddr _ ty => ty + | Eunop _ _ ty => ty + | Ebinop _ _ _ ty => ty + | Efield _ _ ty => ty + | Eindex _ _ ty => ty + | Ecall0 _ ty => ty + | Ecall1 _ _ ty => ty + end. diff --git a/src/backend/Expressions/ExpStacked.v b/src/backend/Expressions/ExpStacked.v new file mode 100755 index 0000000..d119a06 --- /dev/null +++ b/src/backend/Expressions/ExpStacked.v @@ -0,0 +1,102 @@ + +(* Intermediate language after Clabeled. +Merges a stackframe of local environments and return points +into a single stack. +In order to make return point well-defined, +we split function calls into pushing, jumping, and saving return value *) + +Require Import cclib.Coqlib. +Require Import backend.AST. +Require Import backend.Environments.Globalenvs. +Require Import cclib.Integers. +Require Import backend.Ctypes. +Require Import backend.Values.LowValues. +Require Import backend.Options. +Require Import backend.Cop. +Require Import backend.MachineModel. +Require Import backend.MachineModelLow. + +(* This is the first phase where we don't use MiniC expressions. Now +expressions is not a recursive datatype, and intermediate values are +instead stored on the stack. + +Local env references are direct now, and expression values are on the +stack. Expressions that take arguments will pop them off the stack *) +Inductive expr : Type := + | Econst_int256: int256 -> expr (**r 256-bit integer literal *) + | Eglob: ident -> expr (**r variable *) + | Etempvar: nat -> expr (**r temporary variable *) + | Emload: expr (**r memory pointer dereference (unary [*]) *) + | Esload: expr (**r storage pointer dereference (unary [*]) *) + | Eunop: unary_operation -> expr (**r unary operation *) + (* all binary operations use the top of the stack as their first argument and + second on the stack as their second argument. + So when computing the first argument, the stack will have an + unnamed value on top. *) + | Ebinop: binary_operation -> bool -> expr (**r binary operation. bool is a signedness flag *) + | Ecall0: builtin0 -> expr + | Ecall1: builtin1 -> expr. + +(* Describes return types of a function, because this affects how method/function calls are compiled: *) +Inductive ret_type : Type := + | Tvoid_fun: ret_type (* has to push dummy 0 value *) + | Tvoid_method: ret_type (* actually returns nothing *) + | Tvoid_constructor: ret_type (* RETURNs code *) + | Terror: ret_type (* type cannot be returned, like array *) + | Tsome_fun: ret_type (* returns something by pushing onto the stack *) + | Tsome_method: ret_type (* returns with a RETURN command *) + | Tsome_constructor: ret_type. (* pops value, RETURNs code *) + +(* In earlier langauges we had instructions for calling functions. Now we want to compile those into jumps, so we have several types of labels (to enforce uniqueness). + Eventually, these will again be unified into a single lable type. *) +Inductive typed_label : Type := + | Linternal : label -> typed_label (* internal label within the current function *) + | Lcall: label -> typed_label (* function to call *) + | Lreturn: label -> typed_label. (* label for returning to the inside of another function *) + +(* expression arguments are popped off the stack *) +Inductive statement : Type := + | Sskip: statement (* no-op *) + | Srvalue: expr -> statement (* pushes an expression as an rvalue *) + | Slvalue : expr -> statement (* pushes an expression as an lvalue. compiled the same but semantically different. (TODO: there will need to be a future pass where these are the same, and pointers are turned into actual integers.) *) + | Spushvoid : statement (* semantically, pushes a void value onto the stack, as a placeholder *) + | Spop : statement (* remove a value from the stack *) + | Sassign : statement + | Sset : nat -> statement (* nat references a stack position, which must reference a temp in this function's local env. moves the thing on the stack to replace this temp. (So this sets a place in the stack, whereas in previous languages it would have been a temp identifier.) *) + (* cleanup function and return value pushed on stack. + first argument is amount of junk to pop off stack, second is about the return value. *) + | Sdone : nat -> ret_type -> statement + + (* control flow *) + | Spushlabel: typed_label -> statement + | Slabel: label -> statement + | Sjump_call : statement (* unconditional jump to a different function *) + | Sjump_internal: statement (* unconditional jump inside a function. *) + | Sjumpi: statement (* cond, label on stack. jump if nonzero *) + + (* custom ethereum statements *) + | Stransfer : statement (** address.transfer(value), with address & value from the stack. + actually all operands to CALL are on the stack, but some have dummy values. (hopefully, the transfer will eventually get compiled to a single call command.) + pushes success value *) + + (* This no longer takes a list of expressions for arguments, but it still has the number of arguments. + The way this should get compiled in later phases is to take right number of arguments from the stack, put them in memory, do the call, and then load the results from memory to the stack. *) + | Scallmethod : int -> nat -> nat -> statement (* sig, number of args, number of retvals. from stack get addr, val, args *) + + | Slog : nat -> nat -> statement (* number of topics, number of args; both are from stack *) + | Srevert : statement + + (* This fetches arguments from CALLDATA and puts them on the stack. *) + | Sfetchargs : nat -> statement (* how many args to push *) + . + +Definition code : Type := list statement. + +Record function : Type := mkfunction { + fn_code: code; +}. + +Definition genv := Genv.t function type. +(* pair genv with label for start of transaction body *) +Definition program : Type := genv * label. + diff --git a/src/backend/Expressions/SemanticsMiniC.v b/src/backend/Expressions/SemanticsMiniC.v new file mode 100644 index 0000000..e7a813b --- /dev/null +++ b/src/backend/Expressions/SemanticsMiniC.v @@ -0,0 +1,126 @@ +Require Import backend.MachineModel. +Require Import backend.Environments.AllEnvironments. +Require Import backend.Environments.ExtEnv. +Require Import backend.TempModel. +Require Import backend.AST. +Require Import cclib.Coqlib. +Require Import backend.Expressions.ExpMiniC. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.Cop. +Require Import backend.Ctypes. +Require Import backend.Options. +Require Import backend.Values.HighValues. +Require Import backend.AbstractData. + +Section SEMANTICS. + + Variable adata: Type. + + Variable ctx: int256. + + (* machine environment *) + Variable me: machine_env adata. + + (* storage/memory environment *) + Variable ee: ext_env. + + Variable le: temp_env. + + (* Calculate the size of the offset we need, in words. *) + (*Definition sem_field_offset (ty: type) (field: ident) : option int256 := + match ty with + | Tstruct _ fld => + unpack offset, _ <- struct_field fld field ;; Some (Int256.repr (Z.of_nat offset)) + | _ => None + end.*) + + (* rvalues always evaluate to ints. *) + Inductive eval_rvalue: expr -> val -> Prop := + | eval_Econst_int256: forall i ty, + eval_rvalue (Econst_int256 i ty) (Vint i) + | eval_Etempvar: forall id val ty, + PTree.get id le = Some val -> + eval_rvalue (Etempvar id ty) val + | eval_Elvalue: forall e lv v, + (* This case is tricky, it relies on the set of things that eval_lvalue + can evaluate is disjoint from the other ones. + The effect is that any lexpr which is embedded in an rexpr implicitly involves an lookup in ee. *) + eval_lvalue e lv -> + read lv ee = Some v -> + eval_rvalue e v + | eval_Eaddr: forall e lv ty, + eval_lvalue e lv -> + eval_rvalue (Eaddr e ty) (Vptr (LVeid lv)) + | eval_Eunop: forall op a ty v1 v, + eval_rvalue a v1 -> + sem_unary_operation op v1 Tvoid = Some v -> + eval_rvalue (Eunop op a ty) v + | eval_Ebinop: forall op a1 a2 ty1 ty2 ty v1 v2 v, + eval_rvalue a1 v1 -> + eval_rvalue a2 v2 -> + (* For the ethereum backend we mostly ignore types, so to be compatible + with the middle end it is useful to be able to put in arbitrary ty1/ty2. *) + sem_binary_operation op v1 ty1 v2 ty2 = Some v -> + eval_rvalue (Ebinop op a1 a2 ty) v + | eval_Ecall0: forall b ty v, + (me_query me) (Qcall0 b) = v -> + eval_rvalue (Ecall0 b ty) v + | eval_Ecall1: forall b a av ty v, + eval_rvalue a av -> + (me_query me) (Qcall1 b av) = v -> + eval_rvalue (Ecall1 b a ty) v + +(* lvalues always evaluate to extended identifiers. + The `val` type is mutually defined with an `lval` type, + but that is not actually used until in later compiler phases *) +with eval_lvalue: expr -> ident_ext -> Prop := +| eval_Evar: forall id ty lv, + var ctx id ee = Some lv -> + eval_lvalue (Evar id ty) lv +| eval_Eglob: forall id ty lv, + glob id ee = Some lv -> + eval_lvalue (Eglob id ty) lv +| eval_Efield: forall e id ty base lv, + eval_lvalue e base -> + access_field base (typeof e) id = Some lv -> + eval_lvalue (Efield e id ty) lv +| eval_Eindex: forall e1 e2 ty base ind lv, + eval_lvalue e1 base -> + eval_rvalue e2 ind -> + index_array base (typeof e1) ind = Some lv -> + eval_lvalue (Eindex e1 e2 ty) lv +| eval_Ederef: forall e ty v lv, + eval_rvalue e v -> + deref v = Some lv -> + eval_lvalue (Ederef e ty) lv +. + +Inductive eval_rvalues: list expr -> list val -> Prop := + | eval_Enil: + eval_rvalues nil nil + | eval_Econs: forall rv val rvs vals, + eval_rvalue rv val -> + eval_rvalues rvs vals -> + eval_rvalues (rv :: rvs) (val :: vals). + +Remark eval_rvalues_len: forall es vs, + eval_rvalues es vs -> length es = length vs. +Proof. induction es; intros; inv H. auto. simpl. +replace (length vals) with (length es). auto. apply IHes; auto. +Qed. + +(* evaluates to Vunit for None. *) +Inductive eval_optvalue: option expr -> val -> Prop := + | eval_Enone: + eval_optvalue None Vunit + | eval_Esome: forall e v, + eval_rvalue e v -> + eval_optvalue (Some e) v. + +End SEMANTICS. + +Arguments eval_lvalue {adata}. +Arguments eval_rvalue {adata}. +Arguments eval_rvalues {adata}. +Arguments eval_optvalue {adata}. diff --git a/src/backend/GasModel.v b/src/backend/GasModel.v new file mode 100755 index 0000000..d517749 --- /dev/null +++ b/src/backend/GasModel.v @@ -0,0 +1,487 @@ + +Require Import cclib.Coqlib. +Require Import backend.Expressions.ExpStacked. +Require Import backend.Expressions.ExpMiniC. +Require Import backend.Cop. +Require Import backend.Ctypes. +Require Import backend.MachineModel. +Require Import Environments.Globalenvs. + +Local Open Scope nat_scope. + +Ltac REORDER order := assert order as NEWORDER by (intros; omega); rewrite NEWORDER; clear NEWORDER. +Ltac REORDERIN order hyp := assert order as NEWORDER by (intros; omega); rewrite NEWORDER in hyp; clear NEWORDER. + +Opaque Peano.plus. + +(* constants from yellow paper *) +Definition g_base : nat := 2. +Definition g_verylow : nat := 3. +Definition g_low : nat := 5. +Definition g_mid : nat := 8. +Definition g_high : nat := 10. +Definition g_jumpdest : nat := 1. + +Definition g_sload : nat := 200. +Definition g_sset : nat := 20000. (* upper bound SSTORE by assuming it's always zero -> nonzero *) + + +(** Derived constants *) + +Definition gas_mload : nat := g_verylow. +Definition gas_push : nat := g_verylow. +Definition gas_pop : nat := g_base. +Definition gas_calldataload : nat := g_verylow. +Definition gas_codecopy : nat := g_verylow. +Definition gas_dup : nat := g_verylow. +Definition gas_swap : nat := g_verylow. + +(* jumps and labels are inexpensive, +and it's hard to keep track at all levels of the number of each required at each level *) +Definition gas_label : nat := g_jumpdest. +Definition gas_jump : nat := g_mid. +Definition gas_jumpi : nat := g_high. +Definition gas_jump_arg : nat := gas_push + gas_jump. +Definition gas_jumpi_arg : nat := gas_push + gas_jumpi. +(* either fallthrough or jump or conditional jump *) +Definition gas_branch : nat := Nat.max gas_label (Nat.max gas_jump_arg gas_jumpi_arg). + + +Remark gas_jump_le_branch: gas_jump_arg <= gas_branch. +Proof. unfold gas_branch. apply Nat.le_trans with (Nat.max gas_jump_arg gas_jumpi_arg). +apply Nat.le_max_l. apply Nat.le_max_r. +Qed. + +Remark gas_jumpi_le_branch: gas_jumpi_arg <= gas_branch. +Proof. unfold gas_branch. apply Nat.le_trans with (Nat.max gas_jump_arg gas_jumpi_arg). +apply Nat.le_max_r. apply Nat.le_max_r. +Qed. + +Remark gas_label_le_branch: gas_label <= gas_branch. +Proof. unfold gas_branch. apply Nat.le_max_l. +Qed. + +Fixpoint gas_branches (count: nat) : nat := + match count with 0 => 0 | S n => gas_branch + gas_branches n end. + +Lemma gas_branches_split: forall b1 b2, + gas_branches (b1 + b2) = gas_branches b1 + gas_branches b2. +Proof. induction b1; simpl; intros; repeat rewrite Nat.add_0_l. +auto. rewrite Nat.add_succ_l. simpl. rewrite IHb1. omega. +Qed. + +Definition gas_eval_builtin0 (b: builtin0) : nat := + match b with + | Baddress => g_base + | Borigin => g_base + | Bcaller => g_base + | Bcallvalue => g_base + | Bcoinbase => g_base + | Btimestamp => g_base + | Bnumber => g_base + | Bchainid => g_base + | Bselfbalance => g_low + end. + +Definition gas_eval_builtin1 (b: builtin1) : nat := + match b with + (* these are from the yellow paper *) + | Bbalance => 400 + | Bblockhash => 20 + end. + +Definition gas_sha_1 : nat := 30 + 6. (* 30 flat rate plus 6 per word *) +Definition gas_sha_2 : nat := 30 + 12. (* 30 flat rate plus 6 per word *) + +Definition gas_binop (o: binary_operation) : nat := + match o with + | Oadd => g_verylow + | Osub => g_verylow + | Omul => g_low + | Osha_2 => gas_sha_2 + | _ => g_verylow + end. + +Definition gas_unop (o: unary_operation) : nat := + match o with + | Onotbool => g_verylow + | Onotint => g_verylow + | Oneg => gas_push + gas_binop Osub + | Osha_1 => gas_sha_1 + end. + +(* TODO: look up the actual values *) + +(* deref determines whether the expression is automatically dereferenced *) +Fixpoint gas_eval_expr (deref: bool) (e: expr) : nat := + match e with + | Econst_int _ _ => 0 + | Econst_int256 _ _ => gas_push + | Evar _ _ => gas_push + match deref with true => g_sload | false => 0 end + | Eglob _ _ => gas_push + match deref with true => g_sload | false => 0 end + | Etempvar _ _ => gas_dup + | Eaddr _ _ => 0 + | Ederef e _ => (gas_eval_expr deref e) + g_sload + | Eunop o e _ => (gas_eval_expr deref e) + gas_unop o + | Ebinop o e1 e2 _ => (gas_eval_expr deref e1) + (gas_eval_expr deref e2) + gas_binop o + | Efield e _ _ => (gas_eval_expr deref e) + gas_push + gas_sha_2 + g_sload + | Eindex e1 e2 _ => (gas_eval_expr deref e1) + (gas_eval_expr deref e2) + gas_sha_2 + g_sload + | Ecall0 b _ => gas_eval_builtin0 b + | Ecall1 b e _ => gas_eval_expr deref e + gas_eval_builtin1 b + end. + +Fixpoint gas_eval_exprs (deref: bool) (args: list expr) : nat := + match args with + | nil => 0 + | arg::rest => (gas_eval_expr deref arg) + (gas_eval_exprs deref rest) + end. + +Definition gas_eval_optexpr (deref: bool) (optexpr: option expr) : nat := + match optexpr with None => gas_push | Some e => gas_eval_expr deref e end. + +Definition gas_assign (deref: bool) (lv: expr) (rv: expr) (b: nat) : nat := + (gas_eval_expr deref lv) + (gas_eval_expr deref rv) + g_sset + (gas_branches b). + +Definition gas_set_stacked : nat := gas_swap + gas_pop. + +Definition gas_set (deref: bool) (rv: expr) (b: nat) : nat := + (gas_eval_expr deref rv) + gas_set_stacked + (gas_branches b). + +(* TODO: find if gas depends on amount of args and other stuff *) +Definition g_call : nat := 700. +Definition g_callvalue : nat := 9000. + +Definition gas_transfer_base : nat := + g_call + g_callvalue. + +Definition gas_callmethod_base : nat := + g_call + g_callvalue. + +Definition gas_transfer (deref: bool) (a: expr) (v: expr) (b: nat) : nat := + gas_jumpi + gas_push + gas_unop Onotbool + gas_transfer_base + gas_push + gas_push + gas_push + gas_push + gas_push + + (gas_eval_expr deref a) + (gas_eval_expr deref v) + (gas_branches b). + +Definition gas_log_base (ntopics nargs : nat) : nat := + nargs * (2 * g_verylow) (* Write arguments to memory. This is not quite right, because MSTORE can take more than verylow if it extends memory. *) + + 2 * g_verylow (* push memory base/size *) + + 375 * (1 + ntopics). (* the LOG0/.../LOG4 instruction *) + + +Definition gas_log (deref: bool) (topics args : list expr) (b: nat) : nat := + gas_log_base (length topics) (length args) + + (gas_eval_exprs deref topics) + + (gas_eval_exprs deref args) + + (gas_branches b). + +Definition gas_saveretval : nat := gas_set_stacked. + +Definition gas_saveretval' (b: nat) : nat := gas_saveretval + gas_branches b. + +Fixpoint gas_saveretvals (c: nat) : nat := + match c with 0 => 0 | S n => gas_saveretval + gas_saveretvals n end. + +(* read from memory *) +Definition gas_extractretval : nat := 30. +Fixpoint gas_extractretvals (c: nat) : nat := + match c with 0 => 0 | S n => gas_extractretval + gas_extractretvals n end. + +(* put in memory *) +Definition gas_stasharg : nat := 22. +Fixpoint gas_stashargs (c: nat) : nat := + match c with 0 => 0 | S n => gas_stasharg + gas_stashargs n end. + +Definition gas_return_var (deref: bool) (rv: positive) (b: nat) : nat := + (gas_eval_expr deref (Etempvar rv Tvoid)) + (gas_branches b). + +Definition gas_return_none (deref: bool) (b: nat) : nat := + (gas_branches b). + +Definition gas_pops (garbage: nat) : nat := + garbage * gas_pop. + +Definition gas_cleanup (garbage: nat) : nat := + match garbage with + | O => O + | _ => gas_pops garbage + gas_swap + end. + +Definition gas_done_simple (garbage: nat) : nat := + gas_cleanup garbage + 100. + +Definition gas_done (garbage: nat) (b: nat) : nat := + gas_done_simple garbage + gas_branches b. + +Definition gas_return_done (deref: bool) (rv: option positive) (garbage: nat) (b: nat) : nat := + match rv with + | Some rv => gas_return_var deref rv 0 + gas_done garbage 0 + gas_branches b + | None => gas_return_none deref 0 + gas_done garbage 0 + gas_branches b + end. + +(*Lemma gas_return_done_split: forall d r g b1 b2, + gas_return_done d r g (b1 + b2) = gas_done g b1 + gas_return d r b2. +Proof. +intros. unfold gas_return_done. unfold gas_done. unfold gas_return. +rewrite gas_branches_split. simpl. omega. +Qed.*) + +Definition gas_callmethod_stacked (rv_count: nat) (args: nat) : nat := + gas_callmethod_base + + (gas_extractretvals rv_count) + + (gas_stashargs args). + +Definition gas_callmethod (deref: bool) (a: expr) (rv_count: nat) (v: expr) (args: list expr) (b: nat) : nat := + gas_push + gas_push + gas_push + gas_push + gas_push + gas_push + gas_jumpi + gas_unop Onotbool + + gas_eval_expr deref a + + gas_saveretvals rv_count + + gas_eval_expr deref v + + gas_eval_exprs deref args + + gas_callmethod_stacked rv_count (length args) + + gas_branches b. + +Definition gas_call (deref: bool) (args: list expr) (b: nat) : nat := + (gas_eval_exprs deref args) + gas_jump_arg + gas_push + (gas_branches b). + +(* pushing empty temps onto the stack *) +Definition gas_temps (temps: nat) : nat := gas_push * temps. + +(* Push args onto the stack. For internal functions, the args are + already on the stack. We reserve the gas because methods and the + constructor need it, and it's difficult to know in the semantics + whether we're entering a method. So there are three cases: + - Internal function, zero gas. + - Method, gas_calldataload + gas_push. + - Constructor, 3*gas_push + gas_codecopy + gas_mload. + gas_arg is an upper bound of all of them. + *) + +Definition gas_arg_method :nat := gas_calldataload + gas_push. +Definition gas_arg_constr :nat := 3*gas_push + gas_codecopy + gas_mload. + +Definition gas_arg : nat := gas_arg_constr. +Definition gas_args (args: nat) : nat := gas_arg * args. + +Remark gas_arg_bound_method: gas_arg >= gas_arg_method. +Proof. compute. omega. Qed. +Remark gas_arg_bound_constr: gas_arg >= gas_arg_constr. +Proof. compute. omega. Qed. + +Definition gas_intro (temps: nat) (args: nat) : nat := + gas_temps temps + gas_args args. + +Definition gas_callenter (temps: nat) (args: nat) (b: nat) : nat := + (gas_intro temps args) + (gas_branches b). + +Definition gas_eval_cond (deref: bool) (e: expr) (b: nat) : nat := + (gas_eval_expr deref e) + (gas_branches b). + +Definition gas_branching (f: nat -> nat) : Prop := + forall x, f x = f 0 + gas_branches x. + +Ltac gas_branching_auto FUN := unfold gas_branching; intros; simpl; + unfold FUN; simpl; repeat rewrite Nat.add_0_r; auto. + +Remark transfer_branching: forall a b c, gas_branching (gas_transfer a b c). +gas_branching_auto gas_transfer. Qed. + +Remark log_branching: forall a b c, gas_branching (gas_log a b c). +gas_branching_auto gas_log. Qed. + +Remark callmethod_branching: forall a b c d e, gas_branching (gas_callmethod a b c d e). +gas_branching_auto gas_callmethod. Qed. + +Remark callenter_branching: forall a b, gas_branching (gas_callenter a b). +gas_branching_auto gas_callenter. Qed. + +Remark assign_branching: forall a b c, gas_branching (gas_assign a b c). +gas_branching_auto gas_assign. Qed. + +Remark set_branching: forall a b, gas_branching (gas_set a b). +gas_branching_auto gas_set. Qed. + +Remark call_branching: forall a b, gas_branching (gas_call a b). +gas_branching_auto gas_call. Qed. + +Remark cond_branching: forall a b, gas_branching (gas_eval_cond a b). +gas_branching_auto gas_eval_cond. Qed. + +(*Remark return_branching: forall a b, gas_branching (gas_return a b). +gas_branching_auto gas_return. Qed.*) + +Remark done_branching: forall a, gas_branching (gas_done a). +gas_branching_auto gas_done. Qed. + +Remark retval_branching: gas_branching gas_saveretval'. +gas_branching_auto gas_saveretval'. Qed. + + +Lemma succ_pred: forall x, 0 < x -> S (Nat.pred x) = x. +Proof. +destruct x; intros. +omega. +simpl. reflexivity. +Qed. + +Lemma gas_add_sub: forall c g g', + g + c <= g' -> + exists g'', g' = g'' + c /\ g <= g''. +Proof. +induction c; simpl; intros. +exists g'. rewrite Nat.add_0_r in H. split; auto. +rewrite Nat.add_comm in H. rewrite Nat.add_succ_l in H. +rewrite Nat.add_comm in H. rewrite <- Nat.add_succ_l in H. +destruct (IHc (S g) g' H) as (g_0 & g_0_eq & g_0_lt). +exists (pred g_0). +assert (S (Init.Nat.pred g_0) = g_0). +{ rewrite succ_pred; auto. unfold lt. + apply Nat.le_trans with (S g); auto. omega. } +split. + +rewrite Nat.add_comm. rewrite Nat.add_succ_l. +rewrite Nat.add_comm. rewrite <- Nat.add_succ_l. +rewrite H0; auto. +apply le_S_n. rewrite H0. auto. +Qed. + +Lemma le_add: forall g c, + g <= g + c. +Proof. induction c; intros. +omega. apply Nat.le_trans with (g + c); auto. +rewrite Nat.add_succ_r. omega. +Qed. + +Lemma le_pad: forall g c c', + c' <= c -> c' <= c + g. +Proof. intros. apply Nat.le_trans with c; auto. apply le_add. +Qed. + +Lemma le_diff: forall g c c', + c' <= c -> g + c' <= g + c. +Proof. induction g; intros. +repeat rewrite Nat.add_0_l. assumption. +repeat rewrite Nat.add_succ_l. +apply le_n_S. apply IHg. assumption. +Qed. + +Lemma le_diff2: forall g c c', + c' <= c -> c' + g <= c + g. +Proof. intros. +rewrite Nat.add_comm. rewrite (Nat.add_comm c g). +apply le_diff. auto. +Qed. + +Lemma le_pad': forall g c c', + c' + g <= c -> c' <= c. +Proof. intros. apply Nat.le_trans with (c' + g); auto. apply le_pad. auto. +Qed. + +Lemma le_add2: forall g g' c c', + c' <= c -> g' <= g -> c' + g' <= c + g. +Proof. intros. +apply Nat.le_trans with (c' + g). +apply le_diff; auto. +apply le_diff2; auto. +Qed. + +Lemma gas_partial_add_sub: forall c g g', + g + c <= g' -> + forall c', + c' <= c -> + exists g'', g' = g'' + c' /\ g <= g''. +Proof. +intros. apply gas_add_sub. apply Nat.le_trans with (g + c); auto. +apply le_diff. auto. +Qed. + + + +Lemma gas_branch_label: forall g g', + g + gas_branch <= g' -> + exists g'', g' = g'' + gas_label /\ g <= g''. +Proof. +intros. eapply gas_partial_add_sub; eauto. apply gas_label_le_branch. +Qed. + +Lemma gas_branch_jump: forall g g', + g + gas_branch <= g' -> + exists g'', g' = g'' + gas_jump_arg /\ g <= g''. +Proof. +intros. eapply gas_partial_add_sub; eauto. apply gas_jump_le_branch. +Qed. + +Lemma gas_branch_jumpi: forall g g', + g + gas_branch <= g' -> + exists g'', g' = g'' + gas_jumpi_arg /\ g <= g''. +Proof. +intros. eapply gas_partial_add_sub; eauto. (* apply gas_jumpi_le_branch. *) +Qed. + + +Ltac same_gas_used GAS g' := apply gas_add_sub in GAS; destruct GAS as (g'' & g_eq & GAS); subst g'. + + +Lemma gas_split_branches: forall g f, + gas_branching f -> + g + f 2 = g + gas_branch + f 0 + gas_branch. +Proof. +unfold gas_branching. intros. assert (A := H 2). +rewrite A. +simpl. +omega. +Qed. + +Lemma gas_extract_branch: forall g f, + gas_branching f -> + g + f 1 = g + f 0 + gas_branch. +Proof. unfold gas_branching. intros. assert (A := H 1). +rewrite A; simpl; omega. +Qed. + +Ltac gas_use_2 GAS g' g_0 := + same_gas_used GAS g'; rename g'' into g_0; + same_gas_used GAS g_0. + +Ltac gas_use_3 GAS g' g_0 g_1 := + same_gas_used GAS g'; rename g'' into g_0; + same_gas_used GAS g_0; rename g'' into g_1; + same_gas_used GAS g_1. + +Ltac gas_use_4 GAS g' g_0 g_1 g_2 := + same_gas_used GAS g'; rename g'' into g_0; + same_gas_used GAS g_0; rename g'' into g_1; + same_gas_used GAS g_1; rename g'' into g_2; + same_gas_used GAS g_2. + +(* Stacked expression *) + +Definition gas_eval_stacked_expr (e: ExpStacked.expr) : nat := + match e with + | ExpStacked.Econst_int256 _ => gas_push + | ExpStacked.Eglob _ => gas_push + | ExpStacked.Etempvar _ => gas_dup + | ExpStacked.Esload => g_sload + | ExpStacked.Emload => g_sload + | ExpStacked.Eunop o => gas_unop o + | ExpStacked.Ebinop o _ => gas_binop o + | ExpStacked.Ecall0 b => gas_eval_builtin0 b + | ExpStacked.Ecall1 b => gas_eval_builtin1 b + end. + +Fixpoint gas_eval_stacked_exprs (args: list ExpStacked.expr) : nat := + match args with + | nil => 0 + | arg::rest => (gas_eval_stacked_expr arg) + (gas_eval_stacked_exprs rest) + end. + +Definition gas_eval_stacked_optexpr (optexpr: option ExpStacked.expr) : nat := + match optexpr with None => 0 | Some e => gas_eval_stacked_expr e end. + +Definition gas_assign_stacked : nat := g_sset. + +Definition gas_eval_stacked_cond (cond: ExpStacked.expr) (b: nat) : nat := + (gas_eval_stacked_expr cond) + (gas_branches b). + +(* Method multiplexer *) + +(* TODO: make it depend on the number of methods *) +Definition gas_multiplex (methods: nat) : nat := 500. + diff --git a/src/backend/IndexLib.v b/src/backend/IndexLib.v new file mode 100755 index 0000000..4e982a4 --- /dev/null +++ b/src/backend/IndexLib.v @@ -0,0 +1,369 @@ +(* Library for building map index types *) + + +Require Import cclib.Coqlib. +Require Import cclib.Integers. + +Local Open Scope nat_scope. + +(* interleave with zeros *) +Fixpoint sparse_positive (x: positive) : positive := + match x with + | xH => xO xH + | xO y => xO (xO (sparse_positive y)) + | xI y => xO (xI (sparse_positive y)) + end. + +(* interleave with zeros, after the first bit *) +Definition delay_sparse (x: positive) : positive := + match x with + | xH => xH + | xI r => xI (sparse_positive r) + | xO r => xO (sparse_positive r) + end. + +Fixpoint pos_measure (x: positive) : nat := + match x with + | xH => S O + | xI r => S (pos_measure r) + | xO r => S (pos_measure r) + end. + +Fixpoint pick_first (x y: positive) (m: nat) {struct m} : positive := + match m with + | O => xH (* should never reach here *) + | S n => + match x with + | xH => xI (delay_sparse y) + | xO r => xO (pick_first y r n) + | xI r => xI (pick_first y r n) + end + end. + +(* Function to injectively map two positives into +a single positive *) +Definition inject_positive (x y: positive) : positive := + pick_first x y ((pos_measure x) + (pos_measure y)). + +Lemma measure_comm: forall x y, (pos_measure x + pos_measure y) = (pos_measure y + pos_measure x). +Proof. +intros. apply plus_comm. +Qed. + +Lemma measure_pos: forall x y, + ((pos_measure x + S (pos_measure y)) = S (pos_measure x + pos_measure y)). +Proof. +auto. +Qed. + +Lemma measure_plus_one: forall x, pos_measure x + 1 = S (pos_measure x). +Proof. +intro. omega. +Qed. + +Lemma delay_matters: forall x y, delay_sparse x = sparse_positive y -> False. +Proof. +intros x y. revert x. +induction y. + +destruct x. simpl. discriminate. +simpl. injection; clear H. +destruct x; simpl; discriminate. +unfold delay_sparse. simpl. discriminate. + +destruct x; unfold delay_sparse; simpl. +discriminate. +injection; clear H. +destruct x; simpl. +injection; clear H. +fold (delay_sparse (xI x)). +apply IHy. +injection; clear H. +fold (delay_sparse (xO x)). +apply IHy. + +injection; clear H. +fold (delay_sparse xH). +apply IHy. + +discriminate. + +intro x. destruct x; simpl; try discriminate. +injection; clear H. destruct x; simpl; discriminate. +Qed. + +Lemma sparse_is_sparse: forall x y z, + inject_positive x y = sparse_positive z -> False. +Proof. +intros x y z. revert x y. +induction z; intros x y; simpl. + +- + destruct x; unfold inject_positive; simpl. + + discriminate. + + injection; clear H. + + destruct y; simpl; try rewrite measure_pos; simpl. + injection; clear H. + fold (inject_positive x y). + apply IHz. + + discriminate. + + rewrite measure_plus_one. simpl. injection; clear H. + apply delay_matters. + + discriminate. + +- + unfold inject_positive. + destruct x; simpl. + destruct y; rewrite measure_comm; simpl; discriminate. + destruct y; simpl; + try rewrite measure_pos; try rewrite measure_plus_one; simpl. + discriminate. + injection; clear H. + apply IHz. + discriminate. + discriminate. + +- + destruct x; try discriminate. + unfold inject_positive; simpl. + injection; clear H. + destruct y; simpl; + try rewrite measure_pos; try rewrite measure_plus_one; simpl; discriminate. +Qed. + +Lemma delay_sparse_is_sparse: forall x y z, + inject_positive x y = delay_sparse z -> False. +Proof. +intros x y z. +destruct z; unfold delay_sparse; + destruct x; unfold inject_positive; simpl; try discriminate; + injection; clear H. +- + rewrite measure_comm. + fold (inject_positive y x). + apply sparse_is_sparse. +- + apply delay_matters. +- + rewrite measure_comm. + fold (inject_positive y x). + apply sparse_is_sparse. +Qed. + +Lemma injective_sparse: forall x y, sparse_positive x = sparse_positive y -> x = y. +Proof. +induction x; simpl. +- + destruct y; simpl. + intro H. + injection H. + intro. + assert (x=y). + apply IHx. + exact H0. + rewrite H1; reflexivity. + discriminate. + discriminate. +- + destruct y; simpl. + discriminate. + intro H; injection H. + intro. + assert (x=y). + apply IHx; exact H0. + rewrite H1; reflexivity. + discriminate. +- + destruct y; simpl. + discriminate. + discriminate. + intro. reflexivity. +Qed. + +Lemma injective_delay_sparse: forall x y, delay_sparse x = delay_sparse y -> x = y. +Proof. +destruct x; destruct y; simpl; intro; + try discriminate; try reflexivity; injection H; intro. +assert (x=y). +apply injective_sparse. exact H0. +rewrite H1; reflexivity. +assert (x=y). +apply injective_sparse. exact H0. +rewrite H1; reflexivity. +Qed. + + +Lemma injective_positive_1: + forall w x y z, + inject_positive w x = inject_positive y z -> w = y. +Proof. +intros w x y z. +revert x w z. +induction y; intros x w z; unfold inject_positive; simpl. ++ + destruct z; destruct x; destruct w; simpl; + try rewrite ?measure_pos; try rewrite ?measure_plus_one; simpl; + try discriminate; intro; injection H; clear H. + * + fold (inject_positive w x). + fold (inject_positive y z). + intro. + assert (w=y). + revert H. + apply IHy. + rewrite H0. + reflexivity. + * + fold (inject_positive y z). + intro. + exfalso. + apply sparse_is_sparse with y z x. + symmetry. + exact H. + * + fold (inject_positive y z). + intro. + exfalso. + apply delay_sparse_is_sparse with y z w. + symmetry. + exact H. + * + fold (inject_positive w x). + fold (inject_positive y z). + intro. + assert (w=y). + revert H. + apply IHy. + rewrite H0. + reflexivity. + * + fold (inject_positive y z). + intro; exfalso; apply sparse_is_sparse with y z x; symmetry; exact H. + * + fold (inject_positive w x). + intro; exfalso; apply delay_sparse_is_sparse with w x y; exact H. + * + intro; exfalso. + apply delay_matters with y x. + symmetry. exact H. + * + intro. + assert (w=y). + apply injective_delay_sparse. + exact H. + rewrite H0; reflexivity. ++ + destruct z; destruct x; destruct w; simpl; + try rewrite ?measure_pos; try rewrite ?measure_plus_one; simpl; + try discriminate; intro; injection H; clear H. + * + fold (inject_positive w x). + fold (inject_positive y z). + intro. + assert (w=y). + revert H. + apply IHy. + rewrite H0. + reflexivity. + * + fold (inject_positive y z). + intro. exfalso. + apply delay_sparse_is_sparse with y z w. + symmetry. exact H. + * + fold (inject_positive w x). + fold (inject_positive y z). + intro. + assert (w=y). + revert H; apply IHy. + rewrite H0; reflexivity. + * + fold (inject_positive w x). + intro. exfalso. + apply delay_sparse_is_sparse with w x y. + exact H. + * + intro. assert (w=y). + apply injective_delay_sparse. + exact H. + rewrite H0; reflexivity. ++ + destruct w; simpl; try discriminate; try reflexivity. + rewrite measure_comm. + fold (inject_positive x w). + intro H; injection H; clear H. + intro. exfalso. + apply delay_sparse_is_sparse with x w z. + exact H. +Qed. + +Lemma injective_positive_2: + forall w x y z, + inject_positive w x = inject_positive y z -> x = z. +Proof. +intros w x y z. +destruct w; destruct y; + unfold inject_positive; simpl; + try rewrite ?measure_pos; simpl; + try discriminate; + intro H; injection H; clear H; + try rewrite measure_comm; simpl. +- + fold (inject_positive x w). + rewrite measure_comm; fold (inject_positive z y). + apply injective_positive_1. +- + fold (inject_positive x w). + intro; exfalso. + apply delay_sparse_is_sparse with x w z; exact H. +- + fold (inject_positive x w). + rewrite measure_comm; fold (inject_positive z y). + apply injective_positive_1. +- + fold (inject_positive z y). + intro; exfalso. + apply delay_sparse_is_sparse with z y x; symmetry; exact H. +- + apply injective_delay_sparse. +Qed. + + + +Definition int_index (i : int256) : positive := + let (i, _) := i in + match i with + | Z0 => xH + | Zpos p => xO p + | Zneg p => xI p + end. + +Lemma int_index_injective: forall (x y: int256), int_index x = int_index y -> x = y. +Proof. +unfold int_index; destruct x; destruct y; intros. +apply Int256.mkint_eq. +destruct intval; destruct intval0; try (inversion H); try reflexivity. +Qed. + +Definition int32_index (i : int) : positive := + let (i, _) := i in + match i with + | Z0 => xH + | Zpos p => xO p + | Zneg p => xI p + end. + +Lemma int32_index_injective: forall (x y: int), int32_index x = int32_index y -> x = y. +Proof. +unfold int32_index; destruct x; destruct y; intros. +apply Int.mkint_eq. +destruct intval; destruct intval0; try (inversion H); try reflexivity. +Qed. + + diff --git a/src/backend/MachineModel.v b/src/backend/MachineModel.v new file mode 100755 index 0000000..dba5484 --- /dev/null +++ b/src/backend/MachineModel.v @@ -0,0 +1,183 @@ +(* The EVM allows contracts to query lots of things +about itself and about the blockchain. +These are the "builtin" expresions. +We treat these like global variable lookups, because +the only way to access the values at any level is fixed. +For example, to compute caller you need to use the +"caller" builtin expression up in DeepSEA, and this must +compile into a "CALLER" command in the bytecode. +*) + +Require Import cclib.Coqlib. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.IndexLib. +Require Import backend.Values.LowValues. +Require Import backend.Values.HighValues. +Require Import backend.Environments.ExtEnv. +Require Import backend.AbstractData. + +(* Defines the set of builtin expressions in the EVM. + They are split into datatypes based on the number of arguments they take. +*) +Inductive builtin0: Type := +| Baddress: builtin0 +| Borigin: builtin0 +| Bcaller: builtin0 +| Bcallvalue: builtin0 +| Bcoinbase: builtin0 +| Btimestamp: builtin0 +| Bnumber: builtin0 +| Bchainid: builtin0 +| Bselfbalance: builtin0. + +Inductive builtin1: Type := +| Bbalance: builtin1 +| Bblockhash: builtin1. + +Inductive state_query : Type := +| Qcall0: builtin0 -> state_query +| Qcall1: builtin1 -> val -> state_query. + +(* Each language semantics has a machine_env (in fact, the same one) +which specifies the result of doing the EVM builtin operations. + +For builtin queries the machine_env contains the information, and +below we define a relation me_query which presents this in a uniform +way to the rest of the compiler correctness proof. + +The machine_env currently contains two abstract relations for +transfers and methods calls to other contracts. These are mostly +placeholders, they will probably need some redesign to be useful to +write code about contracts. + +The entire backend is parametrized on an abstract data type which +keeps state related to transfers and method calls, to be filled in +at a later time. + *) + +Section WITH_DATA. +Context adata {data_ops: CompatDataOps adata}. + +Record machine_env : Type := mkmachine { + me_address : int256; (* Todo: make it int160. *) + me_origin : int256; (* Todo: make it int160. *) + me_caller : int256; + me_callvalue : int256; + me_coinbase : int256; + me_timestamp : int256; + me_number : int256; + me_chainid : int256; + me_selfbalance : int256; + me_balance : int256 -> int256; (* Todo: make it int160. *) + me_blockhash : int256 -> int256; + + (* todo: this is bad because it doesn't deal with potential reentrancy. *) + me_transfer : forall (addr value: val)(prev_dat new_data: adata) (success: int256), Prop; + (* addr, sig, value, args, prev_data, prev_storage, new_data, new_storage, success, retvals *) + me_callmethod : val -> int -> val -> list val -> adata -> ext_env -> adata -> ext_env -> int256 -> list val -> Prop; + me_log : forall (topics : list val) (args : list val), adata -> adata +}. + +Definition me_query (me : machine_env) (q: state_query) : val := + match q with + | Qcall0 Baddress => Vint (me_address me) + | Qcall0 Borigin => Vint (me_origin me) + | Qcall0 Bcaller => Vint (me_caller me) + | Qcall0 Bcallvalue => Vint (me_callvalue me) + | Qcall0 Bcoinbase => Vint (me_coinbase me) + | Qcall0 Btimestamp => Vint (me_timestamp me) + | Qcall0 Bnumber => Vint (me_number me) + | Qcall0 Bchainid => Vint (me_chainid me) + | Qcall0 Bselfbalance => Vint (me_selfbalance me) + | Qcall1 Bbalance (Vint addr) => Vint (me_balance me addr) + | Qcall1 Bbalance _ => Vunit (* ill-typed query. *) + | Qcall1 Bblockhash (Vint n) => Vint (me_blockhash me n) + | Qcall1 Bblockhash _ => Vunit (* ill-typed query. *) + end. + +(* +(* This is now done with a relation *) + +Module QueryIndexed <: INDEXED_TYPE. + +Definition t := state_query. + +Definition index (q: state_query) : positive := + match q with + | Qsha_1 v => xO (xO (xO (xO (xO (int_index v))))) + | Qsha_2 v1 v2 => xO (xO (xO (xO (xI (inject_positive (int_index v1) (int_index v2)))))) + | Qaddress => xO (xO (xO (xI (xO xH)))) + | Qbalance v => xO (xO (xO (xI (xI (int_index v))))) + | Qorigin => xO (xO (xI (xO (xO xH)))) + | Qcaller => xO (xO (xI (xO (xI xH)))) + | Qcallvalue => xO (xO (xI (xI (xO xH)))) + | Qcalldataload v => xO (xO (xI (xI (xI (int_index v))))) + | Qcalldatasize => xO (xI (xO (xO (xO xH)))) + | Qcodesize => xO (xI (xO (xO (xI xH)))) + | Qgasprice => xO (xI (xO (xI (xO xH)))) + | Qextcodesize v => xO (xI (xO (xI (xI (int_index v))))) + | Qblockhash v => xO (xI (xI (xO (xO (int_index v))))) + | Qcoinbase => xO (xI (xI (xO (xI xH)))) + | Qtimestamp => xO (xI (xI (xI (xO xH)))) + | Qnumber => xO (xI (xI (xI (xI xH)))) + | Qdifficulty => xI (xO (xO (xO (xO xH)))) + | Qgaslimit => xI (xO (xO (xO (xI xH)))) + | Qgas => xI (xO (xO (xO (xI (xO xH))))) + end. + +Lemma index_inj: forall q r, index q = index r -> q = r. +Proof. +destruct q; destruct r; simpl; try discriminate; try reflexivity; + intro H; injection H; clear H; intro H. +- + replace i0 with i; try reflexivity; apply int_index_injective; exact H. +- + replace i1 with i. replace i2 with i0. + reflexivity. + apply int_index_injective. + apply injective_positive_2 with (int_index i) (int_index i1). + exact H. + apply int_index_injective. + apply injective_positive_1 with (int_index i0) (int_index i2). + exact H. +- + replace i0 with i; try reflexivity; apply int_index_injective; exact H. +- + replace i0 with i; try reflexivity; apply int_index_injective; exact H. +- + replace i0 with i; try reflexivity; apply int_index_injective; exact H. +- + replace i0 with i; try reflexivity; apply int_index_injective; exact H. +Qed. + +Lemma eq: forall (x y: state_query), {x = y} + {x <> y}. +Proof. +intros. decide equality; apply Int256.eq_dec. +Qed. + +End QueryIndexed. + +Module StateMap := IMap(QueryIndexed). + +Definition state_env := StateMap.t int256. + +*) + +End WITH_DATA. + +Arguments me_address {adata}. +Arguments me_origin {adata}. +Arguments me_caller {adata}. +Arguments me_callvalue {adata}. +Arguments me_coinbase {adata}. +Arguments me_timestamp {adata}. +Arguments me_number {adata}. +Arguments me_chainid {adata}. +Arguments me_selfbalance {adata}. +Arguments me_balance {adata}. +Arguments me_blockhash {adata}. +Arguments me_transfer {adata}. +Arguments me_callmethod {adata}. +Arguments me_query {adata}. +Arguments me_log {adata}. diff --git a/src/backend/MachineModelLow.v b/src/backend/MachineModelLow.v new file mode 100644 index 0000000..8b1473c --- /dev/null +++ b/src/backend/MachineModelLow.v @@ -0,0 +1,80 @@ +(* +The MachineModelLow is exactly the same as the MachineModel, +except the me_transfer, me_callmethods, and me_log now expects +a low value instead of a high value. + +*) + +Require Import cclib.Coqlib. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.IndexLib. +Require Import backend.Values.LowValues. +Require Import backend.Environments.StackEnv. +Require Import backend.Environments.HashEnv. +Require Import backend.AbstractData. +Require Import backend.MachineModel. + + +Inductive state_query : Type := +| Qcall0: builtin0 -> state_query +| Qcall1: builtin1 -> val -> state_query. + + +Section WITH_DATA. +Context adata {data_ops: CompatDataOps adata}. + +Record machine_env : Type := mkmachine { + me_address : int256; (* Todo: make it int160. *) + me_origin : int256; (* Todo: make it int160. *) + me_caller : int256; + me_callvalue : int256; + me_coinbase : int256; + me_timestamp : int256; + me_number : int256; + me_chainid : int256; + me_selfbalance : int256; + me_balance : int256 -> int256; (* Todo: make it int160. *) + me_blockhash : int256 -> int256; + + (* todo: this is bad because it doesn't deal with potential reentrancy. *) + me_transfer : forall (addr value: val)(prev_dat new_data: adata) (success: int256), Prop; + (* addr, sig, value, args, prev_data, prev_storage, new_data, new_storage, success, retvals *) + me_callmethod : val -> int -> val -> list val -> adata -> stack_env -> hash_env -> adata -> stack_env -> hash_env -> int256 -> list val -> Prop; + me_log : forall (topics : list val) (args : list val), adata -> adata +}. + +Definition me_query (me : machine_env) (q: state_query) : val := + match q with + | Qcall0 Baddress => Vint (me_address me) + | Qcall0 Borigin => Vint (me_origin me) + | Qcall0 Bcaller => Vint (me_caller me) + | Qcall0 Bcallvalue => Vint (me_callvalue me) + | Qcall0 Bcoinbase => Vint (me_coinbase me) + | Qcall0 Btimestamp => Vint (me_timestamp me) + | Qcall0 Bnumber => Vint (me_number me) + | Qcall0 Bchainid => Vint (me_chainid me) + | Qcall0 Bselfbalance => Vint (me_selfbalance me) + | Qcall1 Bbalance (Vint addr) => Vint (me_balance me addr) + | Qcall1 Bbalance _ => Vunit (* ill-typed query. *) + | Qcall1 Bblockhash (Vint n) => Vint (me_blockhash me n) + | Qcall1 Bblockhash _ => Vunit (* ill-typed query. *) + end. + +End WITH_DATA. + +Arguments me_address {adata}. +Arguments me_origin {adata}. +Arguments me_caller {adata}. +Arguments me_callvalue {adata}. +Arguments me_coinbase {adata}. +Arguments me_timestamp {adata}. +Arguments me_number {adata}. +Arguments me_chainid {adata}. +Arguments me_selfbalance {adata}. +Arguments me_balance {adata}. +Arguments me_blockhash {adata}. +Arguments me_transfer {adata}. +Arguments me_callmethod {adata}. +Arguments me_query {adata}. +Arguments me_log {adata}. diff --git a/src/backend/MemoryModel.v b/src/backend/MemoryModel.v new file mode 100755 index 0000000..b6b870d --- /dev/null +++ b/src/backend/MemoryModel.v @@ -0,0 +1,50 @@ + +Require Import cclib.Coqlib. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.AST. +Require Import backend.IndexLib. + +(* For the ethereum backend we do not use memory injections, but we +still define a dummy datatype in order to make the middle-end more +portable between C and Ethereum. *) +Definition meminj : Set := unit. + +(* The type of offsets into structs etc. It will be int in the C backend and int256 in the Ethereum backend. *) +Definition offset_t := int256. + +(* after hashing, the rest of memory is for retvals and args *) +Definition public_retval_base : Z := 96%Z. +(* according to ABI, this is how big each arg/retval is *) +Definition abi_word_size : Z := 32%Z. + +(* move funsig over to left-align the 4 bytes to 256 bytes *) +Definition funsig_aligned (funsig: int) : int256 := + let funsig := Int.unsigned funsig in + Int256.repr (Z.shiftl funsig 224). + +(* locations in memory for CALL args *) +Definition public_funsig_pos (retval_count: nat) : Z := + public_retval_base + (abi_word_size * (Z.of_nat retval_count)). +Definition public_arg_pos (index: nat) (retval_count: nat) : Z := + (public_funsig_pos retval_count) + ((Z.of_nat index) * abi_word_size) + 4%Z. +Definition public_arg_size (arg_count: nat) : Z := + ((Z.of_nat arg_count) * abi_word_size) + 4%Z. +Definition retval_pos (index: nat) : Z := + public_retval_base + (abi_word_size * (Z.of_nat index)). + +(* arguments to CALL *) +Definition argpos (retval_count: nat) : Z := public_funsig_pos retval_count. +Definition arglen (arg_count: nat) : Z := public_arg_size arg_count. +Definition retpos : Z := public_retval_base. +Definition retlen (retval_count: nat) : Z := abi_word_size * (Z.of_nat retval_count). + +(* according to the ABI specification, +Calldata locations *) +Definition call_data_arg_location (argi: nat) : int256 := + Int256.repr (((Z.of_nat argi) * 32%Z) + 4%Z). + +(* add two consts for codecopy *) +Definition bytes_to_fetch : int256 := Int256.repr 32%Z. + +Definition MemoryLocation : int256 := Int256.repr 128%Z. \ No newline at end of file diff --git a/src/backend/Options.v b/src/backend/Options.v new file mode 100755 index 0000000..7093515 --- /dev/null +++ b/src/backend/Options.v @@ -0,0 +1,51 @@ +(* +Many things can fail during a program run, so a lot of helper functions used to define the semantics of the various intermediate languages return options. + +Also, many things can fail in compilation, due to failed assumptions, so most of the compilation functions return optErr (i.e. an option or an error message). + +Both of these are monad, and this file defines suitable monad instances and tactics that work the same way on both. +*) + +Require Export DeepSpec.lib.Monad.Monad. +Require Export DeepSpec.lib.Monad.OptionMonad. +Require DeepSpec.lib.Monad.OptionMonadLaws. +Require Export DeepSpec.lib.Monad.OptErrMonad. +Require DeepSpec.lib.Monad.OptErrMonadLaws. +Require Export DeepSpec.lib.Monad.MonadInv. + +Require Coq.Strings.String. +Export String.StringSyntax. + +(* H is where to bind. a is new var and H1 is new hypothesis for value of a *) +Ltac BindSome H a H1 := + match type of H with + | @eq ?T _ _ => + match (eval compute in T) with + | option _ => apply (@bind_some _ _ OptionMonadLaws.MonadLaws_option MonadInvBind_option) in H; destruct H as (a & H1 & H) + | optErr _ => apply (@bind_some _ _ OptErrMonadLaws.MonadLaws_optErr MonadInvBind_optErr) in H; destruct H as (a & H1 & H) + end +end. + +(* call on a hypothesis of the form Some expr = Some id *) +Ltac SomeSome H id := injection H; clear H; intro; subst id. + +Definition fromOption {A:Type} (ov : option A) (msg : String.string) := + match ov with + | Some v => Success v + | None => Error msg + end. + +Lemma fromOption_Success : forall A (ov : option A) msg v, + fromOption ov msg = Success v -> + ov = Some v. +Proof. + intros. + destruct ov; simpl in *; congruence. +Qed. + +Export MonadNotation. + +Definition optional_filter (A: Type) (a: option A) : list A := + match a with None => nil | Some a' => a' :: nil end. + + diff --git a/src/backend/Smallstep.v b/src/backend/Smallstep.v new file mode 100755 index 0000000..82e8c8c --- /dev/null +++ b/src/backend/Smallstep.v @@ -0,0 +1,1847 @@ +(* This is almost unchanged from CompCert, except it no longer uses events. + + Almost every definition is used, and the proofs have been adapted. +*) + + +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Tools for small-step operational semantics *) + +(** This module defines generic operations and theorems over + the one-step transition relations that are used to specify + operational semantics in small-step style. *) + +Require Import cclib.Coqlib. +Require Import cclib.Integers. + +Set Implicit Arguments. + +(** * Closures of transitions relations *) + +Section CLOSURES. + +Variable genv: Type. +Variable state: Type. + +(** A one-step transition relation has the following signature. + It is parameterized by a global environment, which does not + change during the transition. It relates the initial state + of the transition with its final state. The [trace] parameter + captures the observable events possibly generated during the + transition. *) + +Variable step: genv -> state -> state -> Prop. + +(** No transitions: stuck state *) + +Definition nostep (ge: genv) (s: state) : Prop := + forall s', ~(step ge s s'). + +(** Zero, one or several transitions. Also known as Kleene closure, + or reflexive transitive closure. *) + +Inductive star (ge: genv): state -> state -> Prop := + | star_refl: forall s, + star ge s s + | star_step: forall s1 s2 s3, + step ge s1 s2 -> star ge s2 s3 -> + star ge s1 s3. + +Lemma star_one: + forall ge s1 s2, step ge s1 s2 -> star ge s1 s2. +Proof. + intros. eapply star_step; eauto. apply star_refl. +Qed. + +Lemma star_two: + forall ge s1 s2 s3, + step ge s1 s2 -> step ge s2 s3 -> + star ge s1 s3. +Proof. + intros. eapply star_step; eauto. apply star_one; auto. +Qed. + +Lemma star_three: + forall ge s1 s2 s3 s4, + step ge s1 s2 -> step ge s2 s3 -> step ge s3 s4 -> + star ge s1 s4. +Proof. + intros. eapply star_step; eauto. eapply star_two; eauto. +Qed. + +Lemma star_four: + forall ge s1 s2 s3 s4 s5, + step ge s1 s2 -> step ge s2 s3 -> + step ge s3 s4 -> step ge s4 s5 -> + star ge s1 s5. +Proof. + intros. eapply star_step; eauto. eapply star_three; eauto. +Qed. + +Lemma star_trans: + forall ge s1 s2, star ge s1 s2 -> + forall s3, star ge s2 s3 -> star ge s1 s3. +Proof. + induction 1; intros. + auto. eapply star_step; eauto. +Qed. + +Lemma star_left: + forall ge s1 s2 s3, + step ge s1 s2 -> star ge s2 s3 -> + star ge s1 s3. +Proof star_step. + +Lemma star_right: + forall ge s1 s2 s3, + star ge s1 s2 -> step ge s2 s3 -> + star ge s1 s3. +Proof. + intros. eapply star_trans. eauto. apply star_one. eauto. +Qed. + +(* +Lemma star_E0_ind: + forall ge (P: state -> state -> Prop), + (forall s, P s s) -> + (forall s1 s2 s3, step ge s1 E0 s2 -> P s2 s3 -> P s1 s3) -> + forall s1 s2, star ge s1 E0 s2 -> P s1 s2. +Proof. + intros ge P BASE REC. + assert (forall s1 t s2, star ge s1 t s2 -> t = E0 -> P s1 s2). + induction 1; intros; subst. + auto. + destruct (Eapp_E0_inv _ _ H2). subst. eauto. + eauto. +Qed. +*) +(** One or several transitions. Also known as the transitive closure. *) + +Inductive plus (ge: genv): state -> state -> Prop := + | plus_left: forall s1 s2 s3, + step ge s1 s2 -> star ge s2 s3 -> + plus ge s1 s3. + +Lemma plus_one: + forall ge s1 s2, + step ge s1 s2 -> plus ge s1 s2. +Proof. + intros. econstructor; eauto. apply star_refl. +Qed. + +Lemma plus_two: + forall ge s1 s2 s3, + step ge s1 s2 -> step ge s2 s3 -> + plus ge s1 s3. +Proof. + intros. eapply plus_left; eauto. apply star_one; auto. +Qed. + +Lemma plus_three: + forall ge s1 s2 s3 s4, + step ge s1 s2 -> step ge s2 s3 -> step ge s3 s4 -> + plus ge s1 s4. +Proof. + intros. eapply plus_left; eauto. eapply star_two; eauto. +Qed. + +Lemma plus_four: + forall ge s1 s2 s3 s4 s5, + step ge s1 s2 -> step ge s2 s3 -> + step ge s3 s4 -> step ge s4 s5 -> + plus ge s1 s5. +Proof. + intros. eapply plus_left; eauto. eapply star_three; eauto. +Qed. + +Lemma plus_star: + forall ge s1 s2, plus ge s1 s2 -> star ge s1 s2. +Proof. + intros. inversion H; subst. + eapply star_step; eauto. +Qed. + +Lemma plus_right: + forall ge s1 s2 s3, + star ge s1 s2 -> step ge s2 s3 -> + plus ge s1 s3. +Proof. + intros. inversion H; subst. simpl. apply plus_one. auto. + eapply plus_left; eauto. + eapply star_right; eauto. +Qed. + +Lemma plus_left': + forall ge s1 s2 s3, + step ge s1 s2 -> plus ge s2 s3 -> + plus ge s1 s3. +Proof. + intros. eapply plus_left; eauto. apply plus_star; auto. +Qed. + +Lemma plus_right': + forall ge s1 s2 s3, + plus ge s1 s2 -> step ge s2 s3 -> + plus ge s1 s3. +Proof. + intros. eapply plus_right; eauto. apply plus_star; auto. +Qed. + +Lemma plus_star_trans: + forall ge s1 s2 s3, + plus ge s1 s2 -> star ge s2 s3 -> plus ge s1 s3. +Proof. + intros. inversion H; subst. + econstructor; eauto. eapply star_trans; eauto. +Qed. + +Lemma star_plus_trans: + forall ge s1 s2 s3, + star ge s1 s2 -> plus ge s2 s3 -> plus ge s1 s3. +Proof. + intros. inversion H; subst. + simpl; auto. + econstructor. eauto. eapply star_trans. eauto. + apply plus_star. eauto. +Qed. + +Lemma plus_trans: + forall ge s1 s2 s3, + plus ge s1 s2 -> plus ge s2 s3 -> plus ge s1 s3. +Proof. + intros. eapply plus_star_trans. eauto. apply plus_star. eauto. +Qed. +(* +Lemma plus_inv: + forall ge s1 t s2, + plus ge s1 t s2 -> + step ge s1 t s2 \/ exists s', exists t1, exists t2, step ge s1 t1 s' /\ plus ge s' t2 s2 /\ t = t1 ** t2. +Proof. + intros. inversion H; subst. inversion H1; subst. + left. rewrite E0_right. auto. + right. exists s3; exists t1; exists (t0 ** t3); split. auto. + split. econstructor; eauto. auto. +Qed. + +Lemma star_inv: + forall ge s1 t s2, + star ge s1 t s2 -> + (s2 = s1 /\ t = E0) \/ plus ge s1 t s2. +Proof. + intros. inv H. left; auto. right; econstructor; eauto. +Qed. + +Lemma plus_ind2: + forall ge (P: state -> trace -> state -> Prop), + (forall s1 t s2, step ge s1 t s2 -> P s1 t s2) -> + (forall s1 t1 s2 t2 s3 t, + step ge s1 t1 s2 -> plus ge s2 t2 s3 -> P s2 t2 s3 -> t = t1 ** t2 -> + P s1 t s3) -> + forall s1 t s2, plus ge s1 t s2 -> P s1 t s2. +Proof. + intros ge P BASE IND. + assert (forall s1 t s2, star ge s1 t s2 -> + forall s0 t0, step ge s0 t0 s1 -> + P s0 (t0 ** t) s2). + induction 1; intros. + rewrite E0_right. apply BASE; auto. + eapply IND. eauto. econstructor; eauto. subst t. eapply IHstar; eauto. auto. + + intros. inv H0. eauto. +Qed. + +Lemma plus_E0_ind: + forall ge (P: state -> state -> Prop), + (forall s1 s2 s3, step ge s1 E0 s2 -> star ge s2 E0 s3 -> P s1 s3) -> + forall s1 s2, plus ge s1 E0 s2 -> P s1 s2. +Proof. + intros. inv H0. exploit Eapp_E0_inv; eauto. intros [A B]; subst. eauto. +Qed. + +(** Counted sequences of transitions *) + +Inductive starN (ge: genv): nat -> state -> trace -> state -> Prop := + | starN_refl: forall s, + starN ge O s E0 s + | starN_step: forall n s t t1 s' t2 s'', + step ge s t1 s' -> starN ge n s' t2 s'' -> t = t1 ** t2 -> + starN ge (S n) s t s''. + +Remark starN_star: + forall ge n s t s', starN ge n s t s' -> star ge s t s'. +Proof. + induction 1; econstructor; eauto. +Qed. + +Remark star_starN: + forall ge s t s', star ge s t s' -> exists n, starN ge n s t s'. +Proof. + induction 1. + exists O; constructor. + destruct IHstar as [n P]. exists (S n); econstructor; eauto. +Qed. + +(** Infinitely many transitions *) + +CoInductive forever (ge: genv): state -> traceinf -> Prop := + | forever_intro: forall s1 t s2 T, + step ge s1 t s2 -> forever ge s2 T -> + forever ge s1 (t *** T). + +Lemma star_forever: + forall ge s1 t s2, star ge s1 t s2 -> + forall T, forever ge s2 T -> + forever ge s1 (t *** T). +Proof. + induction 1; intros. simpl. auto. + subst t. rewrite Eappinf_assoc. + econstructor; eauto. +Qed. + +(** An alternate, equivalent definition of [forever] that is useful + for coinductive reasoning. *) + +Variable A: Type. +Variable order: A -> A -> Prop. + +CoInductive forever_N (ge: genv) : A -> state -> traceinf -> Prop := + | forever_N_star: forall s1 t s2 a1 a2 T1 T2, + star ge s1 t s2 -> + order a2 a1 -> + forever_N ge a2 s2 T2 -> + T1 = t *** T2 -> + forever_N ge a1 s1 T1 + | forever_N_plus: forall s1 t s2 a1 a2 T1 T2, + plus ge s1 t s2 -> + forever_N ge a2 s2 T2 -> + T1 = t *** T2 -> + forever_N ge a1 s1 T1. + +Hypothesis order_wf: well_founded order. + +Lemma forever_N_inv: + forall ge a s T, + forever_N ge a s T -> + exists t, exists s', exists a', exists T', + step ge s t s' /\ forever_N ge a' s' T' /\ T = t *** T'. +Proof. + intros ge a0. pattern a0. apply (well_founded_ind order_wf). + intros. inv H0. + (* star case *) + inv H1. + (* no transition *) + change (E0 *** T2) with T2. apply H with a2. auto. auto. + (* at least one transition *) + exists t1; exists s0; exists x; exists (t2 *** T2). + split. auto. split. eapply forever_N_star; eauto. + apply Eappinf_assoc. + (* plus case *) + inv H1. + exists t1; exists s0; exists a2; exists (t2 *** T2). + split. auto. + split. inv H3. auto. + eapply forever_N_plus. econstructor; eauto. eauto. auto. + apply Eappinf_assoc. +Qed. + +Lemma forever_N_forever: + forall ge a s T, forever_N ge a s T -> forever ge s T. +Proof. + cofix COINDHYP; intros. + destruct (forever_N_inv H) as [t [s' [a' [T' [P [Q R]]]]]]. + rewrite R. apply forever_intro with s'. auto. + apply COINDHYP with a'; auto. +Qed. + +(** Yet another alternative definition of [forever]. *) + +CoInductive forever_plus (ge: genv) : state -> traceinf -> Prop := + | forever_plus_intro: forall s1 t s2 T1 T2, + plus ge s1 t s2 -> + forever_plus ge s2 T2 -> + T1 = t *** T2 -> + forever_plus ge s1 T1. + +Lemma forever_plus_inv: + forall ge s T, + forever_plus ge s T -> + exists s', exists t, exists T', + step ge s t s' /\ forever_plus ge s' T' /\ T = t *** T'. +Proof. + intros. inv H. inv H0. exists s0; exists t1; exists (t2 *** T2). + split. auto. + split. exploit star_inv; eauto. intros [[P Q] | R]. + subst. simpl. auto. econstructor; eauto. + traceEq. +Qed. + +Lemma forever_plus_forever: + forall ge s T, forever_plus ge s T -> forever ge s T. +Proof. + cofix COINDHYP; intros. + destruct (forever_plus_inv H) as [s' [t [T' [P [Q R]]]]]. + subst. econstructor; eauto. +Qed. + +(** Infinitely many silent transitions *) + +CoInductive forever_silent (ge: genv): state -> Prop := + | forever_silent_intro: forall s1 s2, + step ge s1 E0 s2 -> forever_silent ge s2 -> + forever_silent ge s1. + +(** An alternate definition. *) + +CoInductive forever_silent_N (ge: genv) : A -> state -> Prop := + | forever_silent_N_star: forall s1 s2 a1 a2, + star ge s1 E0 s2 -> + order a2 a1 -> + forever_silent_N ge a2 s2 -> + forever_silent_N ge a1 s1 + | forever_silent_N_plus: forall s1 s2 a1 a2, + plus ge s1 E0 s2 -> + forever_silent_N ge a2 s2 -> + forever_silent_N ge a1 s1. + +Lemma forever_silent_N_inv: + forall ge a s, + forever_silent_N ge a s -> + exists s', exists a', + step ge s E0 s' /\ forever_silent_N ge a' s'. +Proof. + intros ge a0. pattern a0. apply (well_founded_ind order_wf). + intros. inv H0. + (* star case *) + inv H1. + (* no transition *) + apply H with a2. auto. auto. + (* at least one transition *) + exploit Eapp_E0_inv; eauto. intros [P Q]. subst. + exists s0; exists x. + split. auto. eapply forever_silent_N_star; eauto. + (* plus case *) + inv H1. exploit Eapp_E0_inv; eauto. intros [P Q]. subst. + exists s0; exists a2. + split. auto. inv H3. auto. + eapply forever_silent_N_plus. econstructor; eauto. eauto. +Qed. + +Lemma forever_silent_N_forever: + forall ge a s, forever_silent_N ge a s -> forever_silent ge s. +Proof. + cofix COINDHYP; intros. + destruct (forever_silent_N_inv H) as [s' [a' [P Q]]]. + apply forever_silent_intro with s'. auto. + apply COINDHYP with a'; auto. +Qed. + +(** Infinitely many non-silent transitions *) + +CoInductive forever_reactive (ge: genv): state -> traceinf -> Prop := + | forever_reactive_intro: forall s1 s2 t T, + star ge s1 t s2 -> t <> E0 -> forever_reactive ge s2 T -> + forever_reactive ge s1 (t *** T). + +Lemma star_forever_reactive: + forall ge s1 t s2 T, + star ge s1 t s2 -> forever_reactive ge s2 T -> + forever_reactive ge s1 (t *** T). +Proof. + intros. inv H0. rewrite <- Eappinf_assoc. econstructor. + eapply star_trans; eauto. + red; intro. exploit Eapp_E0_inv; eauto. intros [P Q]. contradiction. + auto. +Qed. + +*) + +End CLOSURES. + +(* + +(** * Transition semantics *) + +(** The general form of a transition semantics. *) + +(** [CompCertX:test-compcert-param-final] We parameterize the semantics +over the type of final return value. For whole programs, this shall be +[int]. *) + +Record semantics (RETVAL: Type) : Type := Semantics_gen { + state: Type; + genvtype: Type; + step : genvtype -> state -> trace -> state -> Prop; + initial_state: state -> Prop; + final_state: state -> RETVAL -> Prop; + globalenv: genvtype; + symbolenv: Senv.t +}. + +(** The form used in earlier CompCert versions, for backward compatibility. *) + +Definition Semantics {state funtype vartype: Type} + {RETVAL: Type} + (step: Genv.t funtype vartype -> state -> trace -> state -> Prop) + (initial_state: state -> Prop) + (final_state: state -> RETVAL -> Prop) + (globalenv: Genv.t funtype vartype) := + {| state := state; + genvtype := Genv.t funtype vartype; + step := step; + initial_state := initial_state; + final_state := final_state; + globalenv := globalenv; + symbolenv := Genv.to_senv globalenv |}. + +(** Handy notations. *) + +Notation " 'Step' L " := (step L (globalenv L)) (at level 1) : smallstep_scope. +Notation " 'Star' L " := (star (step L) (globalenv L)) (at level 1) : smallstep_scope. +Notation " 'Plus' L " := (plus (step L) (globalenv L)) (at level 1) : smallstep_scope. +Notation " 'Forever_silent' L " := (forever_silent (step L) (globalenv L)) (at level 1) : smallstep_scope. +Notation " 'Forever_reactive' L " := (forever_reactive (step L) (globalenv L)) (at level 1) : smallstep_scope. +Notation " 'Nostep' L " := (nostep (step L) (globalenv L)) (at level 1) : smallstep_scope. + +Open Scope smallstep_scope. + +(** * Forward simulations between two transition semantics. *) + +(** The general form of a forward simulation. *) + +Record fsim_properties {RETVAL: Type} (L1 L2: semantics RETVAL) (index: Type) + (order: index -> index -> Prop) + (match_states: index -> state L1 -> state L2 -> Prop) : Prop := { + fsim_order_wf: well_founded order; + fsim_match_initial_states: + forall s1, initial_state L1 s1 -> + exists i, exists s2, initial_state L2 s2 /\ match_states i s1 s2; + fsim_match_final_states: + forall i s1 s2 r, + match_states i s1 s2 -> final_state L1 s1 r -> final_state L2 s2 r; + fsim_simulation: + forall s1 t s1', Step L1 s1 t s1' -> + forall i s2, match_states i s1 s2 -> + exists i', exists s2', + (Plus L2 s2 t s2' \/ (Star L2 s2 t s2' /\ order i' i)) + /\ match_states i' s1' s2'; + fsim_public_preserved: + forall id, Senv.public_symbol (symbolenv L2) id = Senv.public_symbol (symbolenv L1) id + }. + +Arguments fsim_properties {_} _ _ _ _ _. + +Inductive forward_simulation {RETVAL: Type} (L1 L2: semantics RETVAL) : Prop := + Forward_simulation (index: Type) + (order: index -> index -> Prop) + (match_states: index -> state L1 -> state L2 -> Prop) + (props: fsim_properties L1 L2 index order match_states). + +Arguments Forward_simulation {RETVAL L1 L2 index} order match_states props. + +(** An alternate form of the simulation diagram *) + +Lemma fsim_simulation': + forall RETVAL: Type, + forall L1 L2: _ RETVAL, + forall index order match_states, fsim_properties L1 L2 index order match_states -> + forall i s1 t s1', Step L1 s1 t s1' -> + forall s2, match_states i s1 s2 -> + (exists i', exists s2', Plus L2 s2 t s2' /\ match_states i' s1' s2') + \/ (exists i', order i' i /\ t = E0 /\ match_states i' s1' s2). +Proof. + intros. exploit @fsim_simulation; eauto. + intros [i' [s2' [A B]]]. intuition. + left; exists i'; exists s2'; auto. + inv H3. + right; exists i'; auto. + left; exists i'; exists s2'; split; auto. econstructor; eauto. +Qed. + +(** ** Forward simulation diagrams. *) + +(** Various simulation diagrams that imply forward simulation *) + +Section FORWARD_SIMU_DIAGRAMS. + +Context {RETVAL: Type}. +Variable L1: semantics RETVAL. +Variable L2: semantics RETVAL. + +Hypothesis public_preserved: + forall id, Senv.public_symbol (symbolenv L2) id = Senv.public_symbol (symbolenv L1) id. + +Variable match_states: state L1 -> state L2 -> Prop. + +Hypothesis match_initial_states: + forall s1, initial_state L1 s1 -> + exists s2, initial_state L2 s2 /\ match_states s1 s2. + +Hypothesis match_final_states: + forall s1 s2 r, + match_states s1 s2 -> + final_state L1 s1 r -> + final_state L2 s2 r. + +(** Simulation when one transition in the first program + corresponds to zero, one or several transitions in the second program. + However, there is no stuttering: infinitely many transitions + in the source program must correspond to infinitely many + transitions in the second program. *) + +Section SIMULATION_STAR_WF. + +(** [order] is a well-founded ordering associated with states + of the first semantics. Stuttering steps must correspond + to states that decrease w.r.t. [order]. *) + +Variable order: state L1 -> state L1 -> Prop. +Hypothesis order_wf: well_founded order. + +Hypothesis simulation: + forall s1 t s1', Step L1 s1 t s1' -> + forall s2, match_states s1 s2 -> + exists s2', + (Plus L2 s2 t s2' \/ (Star L2 s2 t s2' /\ order s1' s1)) + /\ match_states s1' s2'. + +Lemma forward_simulation_star_wf: forward_simulation L1 L2. +Proof. + apply Forward_simulation with order (fun idx s1 s2 => idx = s1 /\ match_states s1 s2); + constructor. +- auto. +- intros. exploit match_initial_states; eauto. intros [s2 [A B]]. + exists s1; exists s2; auto. +- intros. destruct H. eapply match_final_states; eauto. +- intros. destruct H0. subst i. exploit simulation; eauto. intros [s2' [A B]]. + exists s1'; exists s2'; intuition auto. +- auto. +Qed. + +End SIMULATION_STAR_WF. + +Section SIMULATION_STAR. + +(** We now consider the case where we have a nonnegative integer measure + associated with states of the first semantics. It must decrease when we take + a stuttering step. *) + +Variable measure: state L1 -> nat. + +Hypothesis simulation: + forall s1 t s1', Step L1 s1 t s1' -> + forall s2, match_states s1 s2 -> + (exists s2', Plus L2 s2 t s2' /\ match_states s1' s2') + \/ (measure s1' < measure s1 /\ t = E0 /\ match_states s1' s2)%nat. + +Lemma forward_simulation_star: forward_simulation L1 L2. +Proof. + apply forward_simulation_star_wf with (ltof _ measure). + apply well_founded_ltof. + intros. exploit simulation; eauto. intros [[s2' [A B]] | [A [B C]]]. + exists s2'; auto. + exists s2; split. right; split. rewrite B. apply star_refl. auto. auto. +Qed. + +End SIMULATION_STAR. + +(** Simulation when one transition in the first program corresponds + to one or several transitions in the second program. *) + +Section SIMULATION_PLUS. + +Hypothesis simulation: + forall s1 t s1', Step L1 s1 t s1' -> + forall s2, match_states s1 s2 -> + exists s2', Plus L2 s2 t s2' /\ match_states s1' s2'. + +Lemma forward_simulation_plus: forward_simulation L1 L2. +Proof. + apply forward_simulation_star with (measure := fun _ => O). + intros. exploit simulation; eauto. +Qed. + +End SIMULATION_PLUS. + +(** Lock-step simulation: each transition in the first semantics + corresponds to exactly one transition in the second semantics. *) + +Section SIMULATION_STEP. + +Hypothesis simulation: + forall s1 t s1', Step L1 s1 t s1' -> + forall s2, match_states s1 s2 -> + exists s2', Step L2 s2 t s2' /\ match_states s1' s2'. + +Lemma forward_simulation_step: forward_simulation L1 L2. +Proof. + apply forward_simulation_plus. + intros. exploit simulation; eauto. intros [s2' [A B]]. + exists s2'; split; auto. apply plus_one; auto. +Qed. + +End SIMULATION_STEP. + +(** Simulation when one transition in the first program + corresponds to zero or one transitions in the second program. + However, there is no stuttering: infinitely many transitions + in the source program must correspond to infinitely many + transitions in the second program. *) + +Section SIMULATION_OPT. + +Variable measure: state L1 -> nat. + +Hypothesis simulation: + forall s1 t s1', Step L1 s1 t s1' -> + forall s2, match_states s1 s2 -> + (exists s2', Step L2 s2 t s2' /\ match_states s1' s2') + \/ (measure s1' < measure s1 /\ t = E0 /\ match_states s1' s2)%nat. + +Lemma forward_simulation_opt: forward_simulation L1 L2. +Proof. + apply forward_simulation_star with measure. + intros. exploit simulation; eauto. intros [[s2' [A B]] | [A [B C]]]. + left; exists s2'; split; auto. apply plus_one; auto. + right; auto. +Qed. + +End SIMULATION_OPT. + +End FORWARD_SIMU_DIAGRAMS. + +(** ** Forward simulation of transition sequences *) + +Section SIMULATION_SEQUENCES. + +Context {RETVAL: Type}. +Context (L1 L2: semantics RETVAL). +Context index order match_states (S: fsim_properties L1 L2 index order match_states). + +Lemma simulation_star: + forall s1 t s1', Star L1 s1 t s1' -> + forall i s2, match_states i s1 s2 -> + exists i', exists s2', Star L2 s2 t s2' /\ match_states i' s1' s2'. +Proof. + induction 1; intros. + exists i; exists s2; split; auto. apply star_refl. + exploit @fsim_simulation; eauto. intros [i' [s2' [A B]]]. + exploit IHstar; eauto. intros [i'' [s2'' [C D]]]. + exists i''; exists s2''; split; auto. eapply star_trans; eauto. + intuition auto. apply plus_star; auto. +Qed. + +Lemma simulation_plus: + forall s1 t s1', Plus L1 s1 t s1' -> + forall i s2, match_states i s1 s2 -> + (exists i', exists s2', Plus L2 s2 t s2' /\ match_states i' s1' s2') + \/ (exists i', clos_trans _ order i' i /\ t = E0 /\ match_states i' s1' s2). +Proof. + induction 1 using plus_ind2; intros. +(* base case *) + exploit fsim_simulation'; eauto. intros [A | [i' A]]. + left; auto. + right; exists i'; intuition. +(* inductive case *) + exploit fsim_simulation'; eauto. intros [[i' [s2' [A B]]] | [i' [A [B C]]]]. + exploit simulation_star. apply plus_star; eauto. eauto. + intros [i'' [s2'' [P Q]]]. + left; exists i''; exists s2''; split; auto. eapply plus_star_trans; eauto. + exploit IHplus; eauto. intros [[i'' [s2'' [P Q]]] | [i'' [P [Q R]]]]. + subst. simpl. left; exists i''; exists s2''; auto. + subst. simpl. right; exists i''; intuition auto. + eapply t_trans; eauto. eapply t_step; eauto. +Qed. + +Lemma simulation_forever_silent: + forall i s1 s2, + Forever_silent L1 s1 -> match_states i s1 s2 -> + Forever_silent L2 s2. +Proof. + assert (forall i s1 s2, + Forever_silent L1 s1 -> match_states i s1 s2 -> + forever_silent_N (step L2) order (globalenv L2) i s2). + cofix COINDHYP; intros. + inv H. destruct (fsim_simulation S _ _ _ H1 _ _ H0) as [i' [s2' [A B]]]. + destruct A as [C | [C D]]. + eapply forever_silent_N_plus; eauto. + eapply forever_silent_N_star; eauto. + intros. eapply forever_silent_N_forever; eauto. eapply fsim_order_wf; eauto. +Qed. + +Lemma simulation_forever_reactive: + forall i s1 s2 T, + Forever_reactive L1 s1 T -> match_states i s1 s2 -> + Forever_reactive L2 s2 T. +Proof. + cofix COINDHYP; intros. + inv H. + edestruct simulation_star as [i' [st2' [A B]]]; eauto. + econstructor; eauto. +Qed. + +End SIMULATION_SEQUENCES. + +(** ** Composing two forward simulations *) + +Lemma compose_forward_simulations: + forall {RETVAL: Type}, + forall L1 L2 L3, forward_simulation (RETVAL := RETVAL) L1 L2 -> forward_simulation L2 L3 -> forward_simulation L1 L3. +Proof. + intro RETVAL. + intros L1 L2 L3 S12 S23. + destruct S12 as [index order match_states props]. + destruct S23 as [index' order' match_states' props']. + + set (ff_index := (index' * index)%type). + set (ff_order := lex_ord (clos_trans _ order') order). + set (ff_match_states := fun (i: ff_index) (s1: state L1) (s3: state L3) => + exists s2, match_states (snd i) s1 s2 /\ match_states' (fst i) s2 s3). + apply Forward_simulation with ff_order ff_match_states; constructor. +- (* well founded *) + unfold ff_order. apply wf_lex_ord. apply wf_clos_trans. + eapply fsim_order_wf; eauto. eapply fsim_order_wf; eauto. +- (* initial states *) + intros. exploit (fsim_match_initial_states props); eauto. intros [i [s2 [A B]]]. + exploit (fsim_match_initial_states props'); eauto. intros [i' [s3 [C D]]]. + exists (i', i); exists s3; split; auto. exists s2; auto. +- (* final states *) + intros. destruct H as [s3 [A B]]. + eapply (fsim_match_final_states props'); eauto. + eapply (fsim_match_final_states props); eauto. +- (* simulation *) + intros. destruct H0 as [s3 [A B]]. destruct i as [i2 i1]; simpl in *. + exploit (fsim_simulation' props); eauto. intros [[i1' [s3' [C D]]] | [i1' [C [D E]]]]. ++ (* L2 makes one or several steps. *) + exploit @simulation_plus; eauto. intros [[i2' [s2' [P Q]]] | [i2' [P [Q R]]]]. +* (* L3 makes one or several steps *) + exists (i2', i1'); exists s2'; split. auto. exists s3'; auto. +* (* L3 makes no step *) + exists (i2', i1'); exists s2; split. + right; split. subst t; apply star_refl. red. left. auto. + exists s3'; auto. ++ (* L2 makes no step *) + exists (i2, i1'); exists s2; split. + right; split. subst t; apply star_refl. red. right. auto. + exists s3; auto. +- (* symbols *) + intros. transitivity (Senv.public_symbol (symbolenv L2) id); eapply fsim_public_preserved; eauto. +Qed. + +(** * Receptiveness and determinacy *) + +Definition single_events {RETVAL: Type} (L: semantics RETVAL) : Prop := + forall s t s', Step L s t s' -> (length t <= 1)%nat. + +Record receptive {RETVAL: Type} (L: semantics RETVAL) : Prop := + Receptive { + sr_receptive: forall s t1 s1 t2, + Step L s t1 s1 -> match_traces (symbolenv L) t1 t2 -> exists s2, Step L s t2 s2; + sr_traces: + single_events L + }. + +Record determinate {RETVAL: Type} (L: semantics RETVAL) : Prop := + Determinate { + sd_determ: forall s t1 s1 t2 s2, + Step L s t1 s1 -> Step L s t2 s2 -> + match_traces (symbolenv L) t1 t2 /\ (t1 = t2 -> s1 = s2); + sd_traces: + single_events L; + sd_initial_determ: forall s1 s2, + initial_state L s1 -> initial_state L s2 -> s1 = s2; + sd_final_nostep: forall s r, + final_state L s r -> Nostep L s; + sd_final_determ: forall s r1 r2, + final_state L s r1 -> final_state L s r2 -> r1 = r2 + }. + +Section DETERMINACY. + +Context {RETVAL: Type}. +Variable L: semantics RETVAL. +Hypothesis DET: determinate L. + +Lemma sd_determ_1: + forall s t1 s1 t2 s2, + Step L s t1 s1 -> Step L s t2 s2 -> match_traces (symbolenv L) t1 t2. +Proof. + intros. eapply sd_determ; eauto. +Qed. + +Lemma sd_determ_2: + forall s t s1 s2, + Step L s t s1 -> Step L s t s2 -> s1 = s2. +Proof. + intros. eapply sd_determ; eauto. +Qed. + +Lemma star_determinacy: + forall s t s', Star L s t s' -> + forall s'', Star L s t s'' -> Star L s' E0 s'' \/ Star L s'' E0 s'. +Proof. + induction 1; intros. + auto. + inv H2. + right. eapply star_step; eauto. + exploit sd_determ_1. eexact H. eexact H3. intros MT. + exploit (sd_traces DET). eexact H. intros L1. + exploit (sd_traces DET). eexact H3. intros L2. + assert (t1 = t0 /\ t2 = t3). + destruct t1. inv MT. auto. + destruct t1; simpl in L1; try omegaContradiction. + destruct t0. inv MT. destruct t0; simpl in L2; try omegaContradiction. + simpl in H5. split. congruence. congruence. + destruct H1; subst. + assert (s2 = s4) by (eapply sd_determ_2; eauto). subst s4. + auto. +Qed. + +End DETERMINACY. + +(** * Backward simulations between two transition semantics. *) + +Definition safe {RETVAL: Type} (L: semantics RETVAL) (s: state L) : Prop := + forall s', + Star L s E0 s' -> + (exists r, final_state L s' r) + \/ (exists t, exists s'', Step L s' t s''). + +Lemma star_safe: + forall {RETVAL: Type}, + forall (L: semantics RETVAL) s s', + Star L s E0 s' -> safe L s -> safe L s'. +Proof. + intros; red; intros. apply H0. eapply star_trans; eauto. +Qed. + +(** The general form of a backward simulation. *) + +Record bsim_properties {RETVAL: Type} + (L1 L2: semantics RETVAL) (index: Type) + (order: index -> index -> Prop) + (match_states: index -> state L1 -> state L2 -> Prop) : Prop := { + bsim_order_wf: well_founded order; + bsim_initial_states_exist: + forall s1, initial_state L1 s1 -> exists s2, initial_state L2 s2; + bsim_match_initial_states: + forall s1 s2, initial_state L1 s1 -> initial_state L2 s2 -> + exists i, exists s1', initial_state L1 s1' /\ match_states i s1' s2; + bsim_match_final_states: + forall i s1 s2 r, + match_states i s1 s2 -> safe L1 s1 -> final_state L2 s2 r -> + exists s1', Star L1 s1 E0 s1' /\ final_state L1 s1' r; + bsim_progress: + forall i s1 s2, + match_states i s1 s2 -> safe L1 s1 -> + (exists r, final_state L2 s2 r) \/ + (exists t, exists s2', Step L2 s2 t s2'); + bsim_simulation: + forall s2 t s2', Step L2 s2 t s2' -> + forall i s1, match_states i s1 s2 -> safe L1 s1 -> + exists i', exists s1', + (Plus L1 s1 t s1' \/ (Star L1 s1 t s1' /\ order i' i)) + /\ match_states i' s1' s2'; + bsim_public_preserved: + forall id, Senv.public_symbol (symbolenv L2) id = Senv.public_symbol (symbolenv L1) id + }. + +Arguments bsim_properties {_} _ _ _ _ _. + +Inductive backward_simulation {RETVAL: Type} (L1 L2: semantics RETVAL) : Prop := + Backward_simulation (index: Type) + (order: index -> index -> Prop) + (match_states: index -> state L1 -> state L2 -> Prop) + (props: bsim_properties L1 L2 index order match_states). + +Arguments Backward_simulation {RETVAL L1 L2 index} order match_states props. + +(** An alternate form of the simulation diagram *) + +Lemma bsim_simulation': + forall RETVAL: Type, + forall (L1 L2: _ RETVAL) index order match_states, bsim_properties L1 L2 index order match_states -> + forall i s2 t s2', Step L2 s2 t s2' -> + forall s1, match_states i s1 s2 -> safe L1 s1 -> + (exists i', exists s1', Plus L1 s1 t s1' /\ match_states i' s1' s2') + \/ (exists i', order i' i /\ t = E0 /\ match_states i' s1 s2'). +Proof. + intros. exploit @bsim_simulation; eauto. + intros [i' [s1' [A B]]]. intuition. + left; exists i'; exists s1'; auto. + inv H4. + right; exists i'; auto. + left; exists i'; exists s1'; split; auto. econstructor; eauto. +Qed. + +(** ** Backward simulation diagrams. *) + +(** Various simulation diagrams that imply backward simulation. *) + +Section BACKWARD_SIMU_DIAGRAMS. + +Context {RETVAL: Type}. +Variable L1: semantics RETVAL. +Variable L2: semantics RETVAL. + +Hypothesis public_preserved: + forall id, Senv.public_symbol (symbolenv L2) id = Senv.public_symbol (symbolenv L1) id. + +Variable match_states: state L1 -> state L2 -> Prop. + +Hypothesis initial_states_exist: + forall s1, initial_state L1 s1 -> exists s2, initial_state L2 s2. + +Hypothesis match_initial_states: + forall s1 s2, initial_state L1 s1 -> initial_state L2 s2 -> + exists s1', initial_state L1 s1' /\ match_states s1' s2. + +Hypothesis match_final_states: + forall s1 s2 r, + match_states s1 s2 -> final_state L2 s2 r -> final_state L1 s1 r. + +Hypothesis progress: + forall s1 s2, + match_states s1 s2 -> safe L1 s1 -> + (exists r, final_state L2 s2 r) \/ + (exists t, exists s2', Step L2 s2 t s2'). + +Section BACKWARD_SIMULATION_PLUS. + +Hypothesis simulation: + forall s2 t s2', Step L2 s2 t s2' -> + forall s1, match_states s1 s2 -> safe L1 s1 -> + exists s1', Plus L1 s1 t s1' /\ match_states s1' s2'. + +Lemma backward_simulation_plus: backward_simulation L1 L2. +Proof. + apply Backward_simulation with + (fun (x y: unit) => False) + (fun (i: unit) s1 s2 => match_states s1 s2); + constructor; auto. +- red; intros; constructor; intros. contradiction. +- intros. exists tt; eauto. +- intros. exists s1; split. apply star_refl. eauto. +- intros. exploit simulation; eauto. intros [s1' [A B]]. + exists tt; exists s1'; auto. +Qed. + +End BACKWARD_SIMULATION_PLUS. + +End BACKWARD_SIMU_DIAGRAMS. + +(** ** Backward simulation of transition sequences *) + +Section BACKWARD_SIMULATION_SEQUENCES. + +Context {RETVAL: Type}. +Context (L1 L2: _ RETVAL) index order match_states (S: bsim_properties L1 L2 index order match_states). + +Lemma bsim_E0_star: + forall s2 s2', Star L2 s2 E0 s2' -> + forall i s1, match_states i s1 s2 -> safe L1 s1 -> + exists i', exists s1', Star L1 s1 E0 s1' /\ match_states i' s1' s2'. +Proof. + intros s20 s20' STAR0. pattern s20, s20'. eapply star_E0_ind; eauto. +- (* base case *) + intros. exists i; exists s1; split; auto. apply star_refl. +- (* inductive case *) + intros. exploit @bsim_simulation; eauto. intros [i' [s1' [A B]]]. + assert (Star L1 s0 E0 s1'). intuition. apply plus_star; auto. + exploit H0. eauto. eapply star_safe; eauto. intros [i'' [s1'' [C D]]]. + exists i''; exists s1''; split; auto. eapply star_trans; eauto. +Qed. + +Lemma bsim_safe: + forall i s1 s2, + match_states i s1 s2 -> safe L1 s1 -> safe L2 s2. +Proof. + intros; red; intros. + exploit bsim_E0_star; eauto. intros [i' [s1' [A B]]]. + eapply bsim_progress; eauto. eapply star_safe; eauto. +Qed. + +Lemma bsim_E0_plus: + forall s2 t s2', Plus L2 s2 t s2' -> t = E0 -> + forall i s1, match_states i s1 s2 -> safe L1 s1 -> + (exists i', exists s1', Plus L1 s1 E0 s1' /\ match_states i' s1' s2') + \/ (exists i', clos_trans _ order i' i /\ match_states i' s1 s2'). +Proof. + induction 1 using plus_ind2; intros; subst t. +- (* base case *) + exploit bsim_simulation'; eauto. intros [[i' [s1' [A B]]] | [i' [A [B C]]]]. ++ left; exists i'; exists s1'; auto. ++ right; exists i'; intuition. +- (* inductive case *) + exploit Eapp_E0_inv; eauto. intros [EQ1 EQ2]; subst. + exploit bsim_simulation'; eauto. intros [[i' [s1' [A B]]] | [i' [A [B C]]]]. ++ exploit bsim_E0_star. apply plus_star; eauto. eauto. eapply star_safe; eauto. apply plus_star; auto. + intros [i'' [s1'' [P Q]]]. + left; exists i''; exists s1''; intuition. eapply plus_star_trans; eauto. ++ exploit IHplus; eauto. intros [P | [i'' [P Q]]]. + left; auto. + right; exists i''; intuition. eapply t_trans; eauto. apply t_step; auto. +Qed. + +Lemma star_non_E0_split: + forall s2 t s2', Star L2 s2 t s2' -> (length t = 1)%nat -> + exists s2x, exists s2y, Star L2 s2 E0 s2x /\ Step L2 s2x t s2y /\ Star L2 s2y E0 s2'. +Proof. + induction 1; intros. + simpl in H; discriminate. + subst t. + assert (EITHER: t1 = E0 \/ t2 = E0). + unfold Eapp in H2; rewrite app_length in H2. + destruct t1; auto. destruct t2; auto. simpl in H2; omegaContradiction. + destruct EITHER; subst. + exploit IHstar; eauto. intros [s2x [s2y [A [B C]]]]. + exists s2x; exists s2y; intuition. eapply star_left; eauto. + rewrite E0_right. exists s1; exists s2; intuition. apply star_refl. +Qed. + +End BACKWARD_SIMULATION_SEQUENCES. + +(** ** Composing two backward simulations *) + +Section COMPOSE_BACKWARD_SIMULATIONS. + +Context {RETVAL: Type}. +Variable L1: semantics RETVAL. +Variable L2: semantics RETVAL. +Variable L3: semantics RETVAL. +Hypothesis L3_single_events: single_events L3. +Context index order match_states (S12: bsim_properties L1 L2 index order match_states). +Context index' order' match_states' (S23: bsim_properties L2 L3 index' order' match_states'). + +Let bb_index : Type := (index * index')%type. + +Definition bb_order : bb_index -> bb_index -> Prop := lex_ord (clos_trans _ order) order'. + +Inductive bb_match_states: bb_index -> state L1 -> state L3 -> Prop := + | bb_match_later: forall i1 i2 s1 s3 s2x s2y, + match_states i1 s1 s2x -> Star L2 s2x E0 s2y -> match_states' i2 s2y s3 -> + bb_match_states (i1, i2) s1 s3. + +Lemma bb_match_at: forall i1 i2 s1 s3 s2, + match_states i1 s1 s2 -> match_states' i2 s2 s3 -> + bb_match_states (i1, i2) s1 s3. +Proof. + intros. econstructor; eauto. apply star_refl. +Qed. + +Lemma bb_simulation_base: + forall s3 t s3', Step L3 s3 t s3' -> + forall i1 s1 i2 s2, match_states i1 s1 s2 -> match_states' i2 s2 s3 -> safe L1 s1 -> + exists i', exists s1', + (Plus L1 s1 t s1' \/ (Star L1 s1 t s1' /\ bb_order i' (i1, i2))) + /\ bb_match_states i' s1' s3'. +Proof. + intros. + exploit (bsim_simulation' S23); eauto. eapply bsim_safe; eauto. + intros [ [i2' [s2' [PLUS2 MATCH2]]] | [i2' [ORD2 [EQ MATCH2]]]]. +- (* 1 L2 makes one or several transitions *) + assert (EITHER: t = E0 \/ (length t = 1)%nat). + { exploit L3_single_events; eauto. + destruct t; auto. destruct t; auto. simpl. intros. omegaContradiction. } + destruct EITHER. ++ (* 1.1 these are silent transitions *) + subst t. exploit (bsim_E0_plus S12); eauto. + intros [ [i1' [s1' [PLUS1 MATCH1]]] | [i1' [ORD1 MATCH1]]]. +* (* 1.1.1 L1 makes one or several transitions *) + exists (i1', i2'); exists s1'; split. auto. eapply bb_match_at; eauto. +* (* 1.1.2 L1 makes no transitions *) + exists (i1', i2'); exists s1; split. + right; split. apply star_refl. left; auto. + eapply bb_match_at; eauto. ++ (* 1.2 non-silent transitions *) + exploit @star_non_E0_split. apply plus_star; eauto. auto. + intros [s2x [s2y [P [Q R]]]]. + exploit (bsim_E0_star S12). eexact P. eauto. auto. intros [i1' [s1x [X Y]]]. + exploit (bsim_simulation' S12). eexact Q. eauto. eapply star_safe; eauto. + intros [[i1'' [s1y [U V]]] | [i1'' [U [V W]]]]; try (subst t; discriminate). + exists (i1'', i2'); exists s1y; split. + left. eapply star_plus_trans; eauto. eapply bb_match_later; eauto. +- (* 2. L2 makes no transitions *) + subst. exists (i1, i2'); exists s1; split. + right; split. apply star_refl. right; auto. + eapply bb_match_at; eauto. +Qed. + +Lemma bb_simulation: + forall s3 t s3', Step L3 s3 t s3' -> + forall i s1, bb_match_states i s1 s3 -> safe L1 s1 -> + exists i', exists s1', + (Plus L1 s1 t s1' \/ (Star L1 s1 t s1' /\ bb_order i' i)) + /\ bb_match_states i' s1' s3'. +Proof. + intros. inv H0. + exploit star_inv; eauto. intros [[EQ1 EQ2] | PLUS]. +- (* 1. match at *) + subst. eapply bb_simulation_base; eauto. +- (* 2. match later *) + exploit (bsim_E0_plus S12); eauto. + intros [[i1' [s1' [A B]]] | [i1' [A B]]]. ++ (* 2.1 one or several silent transitions *) + exploit bb_simulation_base. eauto. auto. eexact B. eauto. + eapply star_safe; eauto. eapply plus_star; eauto. + intros [i'' [s1'' [C D]]]. + exists i''; exists s1''; split; auto. + left. eapply plus_star_trans; eauto. + destruct C as [P | [P Q]]. apply plus_star; eauto. eauto. + traceEq. ++ (* 2.2 no silent transition *) + exploit bb_simulation_base. eauto. auto. eexact B. eauto. auto. + intros [i'' [s1'' [C D]]]. + exists i''; exists s1''; split; auto. + intuition. right; intuition. + inv H6. left. eapply t_trans; eauto. left; auto. +Qed. + +End COMPOSE_BACKWARD_SIMULATIONS. + +Lemma compose_backward_simulation: + forall RETVAL: Type, + forall L1 L2 L3: _ RETVAL, + single_events L3 -> backward_simulation L1 L2 -> backward_simulation L2 L3 -> + backward_simulation L1 L3. +Proof. + intro RETVAL. + intros L1 L2 L3 L3single S12 S23. + destruct S12 as [index order match_states props]. + destruct S23 as [index' order' match_states' props']. + apply Backward_simulation with (bb_order order order') (bb_match_states L1 L2 L3 match_states match_states'); + constructor. +- (* well founded *) + unfold bb_order. apply wf_lex_ord. apply wf_clos_trans. eapply bsim_order_wf; eauto. eapply bsim_order_wf; eauto. +- (* initial states exist *) + intros. exploit (bsim_initial_states_exist props); eauto. intros [s2 A]. + eapply (bsim_initial_states_exist props'); eauto. +- (* match initial states *) + intros s1 s3 INIT1 INIT3. + exploit (bsim_initial_states_exist props); eauto. intros [s2 INIT2]. + exploit (bsim_match_initial_states props'); eauto. intros [i2 [s2' [INIT2' M2]]]. + exploit (bsim_match_initial_states props); eauto. intros [i1 [s1' [INIT1' M1]]]. + exists (i1, i2); exists s1'; intuition auto. eapply bb_match_at; eauto. +- (* match final states *) + intros i s1 s3 r MS SAFE FIN. inv MS. + exploit (bsim_match_final_states props'); eauto. + eapply star_safe; eauto. eapply bsim_safe; eauto. + intros [s2' [A B]]. + exploit (bsim_E0_star props). eapply star_trans. eexact H0. eexact A. auto. eauto. auto. + intros [i1' [s1' [C D]]]. + exploit (bsim_match_final_states props); eauto. eapply star_safe; eauto. + intros [s1'' [P Q]]. + exists s1''; split; auto. eapply star_trans; eauto. +- (* progress *) + intros i s1 s3 MS SAFE. inv MS. + eapply (bsim_progress props'). eauto. eapply star_safe; eauto. eapply bsim_safe; eauto. +- (* simulation *) + apply bb_simulation; auto. +- (* symbols *) + intros. transitivity (Senv.public_symbol (symbolenv L2) id); eapply bsim_public_preserved; eauto. +Qed. + +(** ** Converting a forward simulation to a backward simulation *) + +Section FORWARD_TO_BACKWARD. + +Context {RETVAL} (L1 L2: _ RETVAL) index order match_states (FS: fsim_properties L1 L2 index order match_states). +Hypothesis L1_receptive: receptive L1. +Hypothesis L2_determinate: determinate L2. + +(** Exploiting forward simulation *) + +Inductive f2b_transitions: state L1 -> state L2 -> Prop := + | f2b_trans_final: forall s1 s2 s1' r, + Star L1 s1 E0 s1' -> + final_state L1 s1' r -> + final_state L2 s2 r -> + f2b_transitions s1 s2 + | f2b_trans_step: forall s1 s2 s1' t s1'' s2' i' i'', + Star L1 s1 E0 s1' -> + Step L1 s1' t s1'' -> + Plus L2 s2 t s2' -> + match_states i' s1' s2 -> + match_states i'' s1'' s2' -> + f2b_transitions s1 s2. + +Lemma f2b_progress: + forall i s1 s2, match_states i s1 s2 -> safe L1 s1 -> f2b_transitions s1 s2. +Proof. + intros i0; pattern i0. apply well_founded_ind with (R := order). + eapply fsim_order_wf; eauto. + intros i REC s1 s2 MATCH SAFE. + destruct (SAFE s1) as [[r FINAL] | [t [s1' STEP1]]]. apply star_refl. +- (* final state reached *) + eapply f2b_trans_final; eauto. + apply star_refl. + eapply fsim_match_final_states; eauto. +- (* L1 can make one step *) + exploit (fsim_simulation FS); eauto. intros [i' [s2' [A MATCH']]]. + assert (B: Plus L2 s2 t s2' \/ (s2' = s2 /\ t = E0 /\ order i' i)). + intuition auto. + destruct (star_inv H0); intuition auto. + clear A. destruct B as [PLUS2 | [EQ1 [EQ2 ORDER]]]. ++ eapply f2b_trans_step; eauto. apply star_refl. ++ subst. exploit REC; eauto. eapply star_safe; eauto. apply star_one; auto. + intros TRANS; inv TRANS. +* eapply f2b_trans_final; eauto. eapply star_left; eauto. +* eapply f2b_trans_step; eauto. eapply star_left; eauto. +Qed. + +Lemma fsim_simulation_not_E0: + forall s1 t s1', Step L1 s1 t s1' -> t <> E0 -> + forall i s2, match_states i s1 s2 -> + exists i', exists s2', Plus L2 s2 t s2' /\ match_states i' s1' s2'. +Proof. + intros. exploit (fsim_simulation FS); eauto. intros [i' [s2' [A B]]]. + exists i'; exists s2'; split; auto. + destruct A. auto. destruct H2. exploit star_inv; eauto. intros [[EQ1 EQ2] | P]; auto. + congruence. +Qed. + +(** Exploiting determinacy *) + +Remark silent_or_not_silent: + forall t, t = E0 \/ t <> E0. +Proof. + intros; unfold E0; destruct t; auto; right; congruence. +Qed. + +Remark not_silent_length: + forall t1 t2, (length (t1 ** t2) <= 1)%nat -> t1 = E0 \/ t2 = E0. +Proof. + unfold Eapp, E0; intros. rewrite app_length in H. + destruct t1; destruct t2; auto. simpl in H. omegaContradiction. +Qed. + +Lemma f2b_determinacy_inv: + forall s2 t' s2' t'' s2'', + Step L2 s2 t' s2' -> Step L2 s2 t'' s2'' -> + (t' = E0 /\ t'' = E0 /\ s2' = s2'') + \/ (t' <> E0 /\ t'' <> E0 /\ match_traces (symbolenv L1) t' t''). +Proof. + intros. + assert (match_traces (symbolenv L2) t' t''). + eapply sd_determ_1; eauto. + destruct (silent_or_not_silent t'). + subst. inv H1. + left; intuition. eapply sd_determ; eauto. + destruct (silent_or_not_silent t''). + subst. inv H1. elim H2; auto. + right; intuition. + eapply match_traces_preserved with (ge1 := (symbolenv L2)); auto. + intros; symmetry; apply (fsim_public_preserved FS). +Qed. + +Lemma f2b_determinacy_star: + forall s s1, Star L2 s E0 s1 -> + forall t s2 s3, + Step L2 s1 t s2 -> t <> E0 -> + Star L2 s t s3 -> + Star L2 s1 t s3. +Proof. + intros s0 s01 ST0. pattern s0, s01. eapply star_E0_ind; eauto. + intros. inv H3. congruence. + exploit f2b_determinacy_inv. eexact H. eexact H4. + intros [[EQ1 [EQ2 EQ3]] | [NEQ1 [NEQ2 MT]]]. + subst. simpl in *. eauto. + congruence. +Qed. + +(** Orders *) + +Inductive f2b_index : Type := + | F2BI_before (n: nat) + | F2BI_after (n: nat). + +Inductive f2b_order: f2b_index -> f2b_index -> Prop := + | f2b_order_before: forall n n', + (n' < n)%nat -> + f2b_order (F2BI_before n') (F2BI_before n) + | f2b_order_after: forall n n', + (n' < n)%nat -> + f2b_order (F2BI_after n') (F2BI_after n) + | f2b_order_switch: forall n n', + f2b_order (F2BI_before n') (F2BI_after n). + +Lemma wf_f2b_order: + well_founded f2b_order. +Proof. + assert (ACC1: forall n, Acc f2b_order (F2BI_before n)). + intros n0; pattern n0; apply lt_wf_ind; intros. + constructor; intros. inv H0. auto. + assert (ACC2: forall n, Acc f2b_order (F2BI_after n)). + intros n0; pattern n0; apply lt_wf_ind; intros. + constructor; intros. inv H0. auto. auto. + red; intros. destruct a; auto. +Qed. + +(** Constructing the backward simulation *) + +Inductive f2b_match_states: f2b_index -> state L1 -> state L2 -> Prop := + | f2b_match_at: forall i s1 s2, + match_states i s1 s2 -> + f2b_match_states (F2BI_after O) s1 s2 + | f2b_match_before: forall s1 t s1' s2b s2 n s2a i, + Step L1 s1 t s1' -> t <> E0 -> + Star L2 s2b E0 s2 -> + starN (step L2) (globalenv L2) n s2 t s2a -> + match_states i s1 s2b -> + f2b_match_states (F2BI_before n) s1 s2 + | f2b_match_after: forall n s2 s2a s1 i, + starN (step L2) (globalenv L2) (S n) s2 E0 s2a -> + match_states i s1 s2a -> + f2b_match_states (F2BI_after (S n)) s1 s2. + +Remark f2b_match_after': + forall n s2 s2a s1 i, + starN (step L2) (globalenv L2) n s2 E0 s2a -> + match_states i s1 s2a -> + f2b_match_states (F2BI_after n) s1 s2. +Proof. + intros. inv H. + econstructor; eauto. + econstructor; eauto. econstructor; eauto. +Qed. + +(** Backward simulation of L2 steps *) + +Lemma f2b_simulation_step: + forall s2 t s2', Step L2 s2 t s2' -> + forall i s1, f2b_match_states i s1 s2 -> safe L1 s1 -> + exists i', exists s1', + (Plus L1 s1 t s1' \/ (Star L1 s1 t s1' /\ f2b_order i' i)) + /\ f2b_match_states i' s1' s2'. +Proof. + intros s2 t s2' STEP2 i s1 MATCH SAFE. + inv MATCH. +- (* 1. At matching states *) + exploit f2b_progress; eauto. intros TRANS; inv TRANS. ++ (* 1.1 L1 can reach final state and L2 is at final state: impossible! *) + exploit (sd_final_nostep L2_determinate); eauto. contradiction. ++ (* 1.2 L1 can make 0 or several steps; L2 can make 1 or several matching steps. *) + inv H2. + exploit f2b_determinacy_inv. eexact H5. eexact STEP2. + intros [[EQ1 [EQ2 EQ3]] | [NOT1 [NOT2 MT]]]. +* (* 1.2.1 L2 makes a silent transition *) + destruct (silent_or_not_silent t2). + (* 1.2.1.1 L1 makes a silent transition too: perform transition now and go to "after" state *) + subst. simpl in *. destruct (star_starN H6) as [n STEPS2]. + exists (F2BI_after n); exists s1''; split. + left. eapply plus_right; eauto. + eapply f2b_match_after'; eauto. + (* 1.2.1.2 L1 makes a non-silent transition: keep it for later and go to "before" state *) + subst. simpl in *. destruct (star_starN H6) as [n STEPS2]. + exists (F2BI_before n); exists s1'; split. + right; split. auto. constructor. + econstructor. eauto. auto. apply star_one; eauto. eauto. eauto. +* (* 1.2.2 L2 makes a non-silent transition, and so does L1 *) + exploit not_silent_length. eapply (sr_traces L1_receptive); eauto. intros [EQ | EQ]. + congruence. + subst t2. rewrite E0_right in H1. + (* Use receptiveness to equate the traces *) + exploit (sr_receptive L1_receptive); eauto. intros [s1''' STEP1]. + exploit fsim_simulation_not_E0. eexact STEP1. auto. eauto. + intros [i''' [s2''' [P Q]]]. inv P. + (* Exploit determinacy *) + exploit not_silent_length. eapply (sr_traces L1_receptive); eauto. intros [EQ | EQ]. + subst t0. simpl in *. exploit @sd_determ_1. eauto. eexact STEP2. eexact H2. + intros. elim NOT2. inv H8. auto. + subst t2. rewrite E0_right in *. + assert (s4 = s2'). eapply sd_determ_2; eauto. subst s4. + (* Perform transition now and go to "after" state *) + destruct (star_starN H7) as [n STEPS2]. exists (F2BI_after n); exists s1'''; split. + left. eapply plus_right; eauto. + eapply f2b_match_after'; eauto. + +- (* 2. Before *) + inv H2. congruence. + exploit f2b_determinacy_inv. eexact H4. eexact STEP2. + intros [[EQ1 [EQ2 EQ3]] | [NOT1 [NOT2 MT]]]. ++ (* 2.1 L2 makes a silent transition: remain in "before" state *) + subst. simpl in *. exists (F2BI_before n0); exists s1; split. + right; split. apply star_refl. constructor. omega. + econstructor; eauto. eapply star_right; eauto. ++ (* 2.2 L2 make a non-silent transition *) + exploit not_silent_length. eapply (sr_traces L1_receptive); eauto. intros [EQ | EQ]. + congruence. + subst. rewrite E0_right in *. + (* Use receptiveness to equate the traces *) + exploit (sr_receptive L1_receptive); eauto. intros [s1''' STEP1]. + exploit fsim_simulation_not_E0. eexact STEP1. auto. eauto. + intros [i''' [s2''' [P Q]]]. + (* Exploit determinacy *) + exploit f2b_determinacy_star. eauto. eexact STEP2. auto. apply plus_star; eauto. + intro R. inv R. congruence. + exploit not_silent_length. eapply (sr_traces L1_receptive); eauto. intros [EQ | EQ]. + subst. simpl in *. exploit @sd_determ_1. eauto. eexact STEP2. eexact H2. + intros. elim NOT2. inv H7; auto. + subst. rewrite E0_right in *. + assert (s3 = s2'). eapply sd_determ_2; eauto. subst s3. + (* Perform transition now and go to "after" state *) + destruct (star_starN H6) as [n STEPS2]. exists (F2BI_after n); exists s1'''; split. + left. apply plus_one; auto. + eapply f2b_match_after'; eauto. + +- (* 3. After *) + inv H. exploit Eapp_E0_inv; eauto. intros [EQ1 EQ2]; subst. + exploit f2b_determinacy_inv. eexact H2. eexact STEP2. + intros [[EQ1 [EQ2 EQ3]] | [NOT1 [NOT2 MT]]]. + subst. exists (F2BI_after n); exists s1; split. + right; split. apply star_refl. constructor; omega. + eapply f2b_match_after'; eauto. + congruence. +Qed. + +End FORWARD_TO_BACKWARD. + +(** The backward simulation *) + +Lemma forward_to_backward_simulation: + forall {RETVAL: Type}, + forall L1 L2: _ RETVAL, + forward_simulation L1 L2 -> receptive L1 -> determinate L2 -> + backward_simulation L1 L2. +Proof. + intro RETVAL. + intros L1 L2 FS L1_receptive L2_determinate. + destruct FS as [index order match_states FS]. + apply Backward_simulation with f2b_order (f2b_match_states L1 L2 match_states); constructor. +- (* well founded *) + apply wf_f2b_order. +- (* initial states exist *) + intros. exploit (fsim_match_initial_states FS); eauto. intros [i [s2 [A B]]]. + exists s2; auto. +- (* initial states *) + intros. exploit (fsim_match_initial_states FS); eauto. intros [i [s2' [A B]]]. + assert (s2 = s2') by (eapply sd_initial_determ; eauto). subst s2'. + exists (F2BI_after O); exists s1; split; auto. econstructor; eauto. +- (* final states *) + intros. inv H. + exploit @f2b_progress; eauto. intros TRANS; inv TRANS. + assert (r0 = r) by (eapply (sd_final_determ L2_determinate); eauto). subst r0. + exists s1'; auto. + inv H4. exploit (sd_final_nostep L2_determinate); eauto. contradiction. + inv H5. congruence. exploit (sd_final_nostep L2_determinate); eauto. contradiction. + inv H2. exploit (sd_final_nostep L2_determinate); eauto. contradiction. +- (* progress *) + intros. inv H. + exploit @f2b_progress; eauto. intros TRANS; inv TRANS. + left; exists r; auto. + inv H3. right; econstructor; econstructor; eauto. + inv H4. congruence. right; econstructor; econstructor; eauto. + inv H1. right; econstructor; econstructor; eauto. +- (* simulation *) + eapply f2b_simulation_step; eauto. +- (* symbols preserved *) + exact (fsim_public_preserved FS). +Qed. + +(** * Transforming a semantics into a single-event, equivalent semantics *) + +Definition well_behaved_traces {RETVAL: Type} (L: semantics RETVAL) : Prop := + forall s t s', Step L s t s' -> + match t with nil => True | ev :: t' => output_trace t' end. + +Section ATOMIC. + +Context {RETVAL: Type}. +Variable L: semantics RETVAL. + +Hypothesis Lwb: well_behaved_traces L. + +Inductive atomic_step (ge: genvtype L): (trace * state L) -> trace -> (trace * state L) -> Prop := + | atomic_step_silent: forall s s', + Step L s E0 s' -> + atomic_step ge (E0, s) E0 (E0, s') + | atomic_step_start: forall s ev t s', + Step L s (ev :: t) s' -> + atomic_step ge (E0, s) (ev :: nil) (t, s') + | atomic_step_continue: forall ev t s, + output_trace (ev :: t) -> + atomic_step ge (ev :: t, s) (ev :: nil) (t, s). + +Definition atomic : semantics RETVAL := {| + state := (trace * state L)%type; + genvtype := genvtype L; + step := atomic_step; + initial_state := fun s => initial_state L (snd s) /\ fst s = E0; + final_state := fun s r => final_state L (snd s) r /\ fst s = E0; + globalenv := globalenv L; + symbolenv := symbolenv L +|}. + +End ATOMIC. + +(** A forward simulation from a semantics [L1] to a single-event semantics [L2] + can be "factored" into a forward simulation from [atomic L1] to [L2]. *) + +Section FACTOR_FORWARD_SIMULATION. + +Context {RETVAL: Type}. +Variable L1: semantics RETVAL. +Variable L2: semantics RETVAL. +Context index order match_states (sim: fsim_properties L1 L2 index order match_states). +Hypothesis L2single: single_events L2. + +Inductive ffs_match: index -> (trace * state L1) -> state L2 -> Prop := + | ffs_match_at: forall i s1 s2, + match_states i s1 s2 -> + ffs_match i (E0, s1) s2 + | ffs_match_buffer: forall i ev t s1 s2 s2', + Star L2 s2 (ev :: t) s2' -> match_states i s1 s2' -> + ffs_match i (ev :: t, s1) s2. + +Lemma star_non_E0_split': + forall s2 t s2', Star L2 s2 t s2' -> + match t with + | nil => True + | ev :: t' => exists s2x, Plus L2 s2 (ev :: nil) s2x /\ Star L2 s2x t' s2' + end. +Proof. + induction 1. simpl. auto. + exploit L2single; eauto. intros LEN. + destruct t1. simpl in *. subst. destruct t2. auto. + destruct IHstar as [s2x [A B]]. exists s2x; split; auto. + eapply plus_left. eauto. apply plus_star; eauto. auto. + destruct t1. simpl in *. subst t. exists s2; split; auto. apply plus_one; auto. + simpl in LEN. omegaContradiction. +Qed. + +Lemma ffs_simulation: + forall s1 t s1', Step (atomic L1) s1 t s1' -> + forall i s2, ffs_match i s1 s2 -> + exists i', exists s2', + (Plus L2 s2 t s2' \/ (Star L2 s2 t s2') /\ order i' i) + /\ ffs_match i' s1' s2'. +Proof. + induction 1; intros. +- (* silent step *) + inv H0. + exploit (fsim_simulation sim); eauto. + intros [i' [s2' [A B]]]. + exists i'; exists s2'; split. auto. constructor; auto. +- (* start step *) + inv H0. + exploit (fsim_simulation sim); eauto. + intros [i' [s2' [A B]]]. + destruct t as [ | ev' t]. ++ (* single event *) + exists i'; exists s2'; split. auto. constructor; auto. ++ (* multiple events *) + assert (C: Star L2 s2 (ev :: ev' :: t) s2'). intuition. apply plus_star; auto. + exploit star_non_E0_split'. eauto. simpl. intros [s2x [P Q]]. + exists i'; exists s2x; split. auto. econstructor; eauto. +- (* continue step *) + inv H0. + exploit star_non_E0_split'. eauto. simpl. intros [s2x [P Q]]. + destruct t. + exists i; exists s2'; split. left. eapply plus_star_trans; eauto. constructor; auto. + exists i; exists s2x; split. auto. econstructor; eauto. +Qed. + +End FACTOR_FORWARD_SIMULATION. + +Theorem factor_forward_simulation: + forall {RETVAL: Type}, + forall L1 L2: _ RETVAL, + forward_simulation L1 L2 -> single_events L2 -> + forward_simulation (atomic L1) L2. +Proof. + intro RETVAL. + intros L1 L2 FS L2single. + destruct FS as [index order match_states sim]. + apply Forward_simulation with order (ffs_match L1 L2 match_states); constructor. +- (* wf *) + eapply fsim_order_wf; eauto. +- (* initial states *) + intros. destruct s1 as [t1 s1]. simpl in H. destruct H. subst. + exploit (fsim_match_initial_states sim); eauto. intros [i [s2 [A B]]]. + exists i; exists s2; split; auto. constructor; auto. +- (* final states *) + intros. destruct s1 as [t1 s1]. simpl in H0; destruct H0; subst. inv H. + eapply (fsim_match_final_states sim); eauto. +- (* simulation *) + eapply ffs_simulation; eauto. +- (* symbols preserved *) + simpl. exact (fsim_public_preserved sim). +Qed. + +(** Likewise, a backward simulation from a single-event semantics [L1] to a semantics [L2] + can be "factored" as a backward simulation from [L1] to [atomic L2]. *) + +Section FACTOR_BACKWARD_SIMULATION. + +Context {RETVAL: Type}. +Variable L1: semantics RETVAL. +Variable L2: semantics RETVAL. +Context index order match_states (sim: bsim_properties L1 L2 index order match_states). +Hypothesis L1single: single_events L1. +Hypothesis L2wb: well_behaved_traces L2. + +Inductive fbs_match: index -> state L1 -> (trace * state L2) -> Prop := + | fbs_match_intro: forall i s1 t s2 s1', + Star L1 s1 t s1' -> match_states i s1' s2 -> + t = E0 \/ output_trace t -> + fbs_match i s1 (t, s2). + +Lemma fbs_simulation: + forall s2 t s2', Step (atomic L2) s2 t s2' -> + forall i s1, fbs_match i s1 s2 -> safe L1 s1 -> + exists i', exists s1', + (Plus L1 s1 t s1' \/ (Star L1 s1 t s1' /\ order i' i)) + /\ fbs_match i' s1' s2'. +Proof. + induction 1; intros. +- (* silent step *) + inv H0. + exploit (bsim_simulation sim); eauto. eapply star_safe; eauto. + intros [i' [s1'' [A B]]]. + exists i'; exists s1''; split. + destruct A as [P | [P Q]]. left. eapply star_plus_trans; eauto. right; split; auto. eapply star_trans; eauto. + econstructor. apply star_refl. auto. auto. +- (* start step *) + inv H0. + exploit (bsim_simulation sim); eauto. eapply star_safe; eauto. + intros [i' [s1'' [A B]]]. + assert (C: Star L1 s1 (ev :: t) s1''). + eapply star_trans. eauto. destruct A as [P | [P Q]]. apply plus_star; eauto. eauto. auto. + exploit @star_non_E0_split'; eauto. simpl. intros [s1x [P Q]]. + exists i'; exists s1x; split. + left; auto. + econstructor; eauto. + exploit L2wb; eauto. +- (* continue step *) + inv H0. unfold E0 in H8; destruct H8; try congruence. + exploit @star_non_E0_split'; eauto. simpl. intros [s1x [P Q]]. + exists i; exists s1x; split. left; auto. econstructor; eauto. simpl in H0; tauto. +Qed. + +Lemma fbs_progress: + forall i s1 s2, + fbs_match i s1 s2 -> safe L1 s1 -> + (exists r, final_state (atomic L2) s2 r) \/ + (exists t, exists s2', Step (atomic L2) s2 t s2'). +Proof. + intros. inv H. destruct t. +- (* 1. no buffered events *) + exploit (bsim_progress sim); eauto. eapply star_safe; eauto. + intros [[r A] | [t [s2' A]]]. ++ (* final state *) + left; exists r; simpl; auto. ++ (* L2 can step *) + destruct t. + right; exists E0; exists (nil, s2'). constructor. auto. + right; exists (e :: nil); exists (t, s2'). constructor. auto. +- (* 2. some buffered events *) + unfold E0 in H3; destruct H3. congruence. + right; exists (e :: nil); exists (t, s3). constructor. auto. +Qed. + +End FACTOR_BACKWARD_SIMULATION. + +Theorem factor_backward_simulation: + forall {RETVAL: Type}, + forall L1 L2: _ RETVAL, + backward_simulation L1 L2 -> single_events L1 -> well_behaved_traces L2 -> + backward_simulation L1 (atomic L2). +Proof. + intro RETVAL. + intros L1 L2 BS L1single L2wb. + destruct BS as [index order match_states sim]. + apply Backward_simulation with order (fbs_match L1 L2 match_states); constructor. +- (* wf *) + eapply bsim_order_wf; eauto. +- (* initial states exist *) + intros. exploit (bsim_initial_states_exist sim); eauto. intros [s2 A]. + exists (E0, s2). simpl; auto. +- (* initial states match *) + intros. destruct s2 as [t s2]; simpl in H0; destruct H0; subst. + exploit (bsim_match_initial_states sim); eauto. intros [i [s1' [A B]]]. + exists i; exists s1'; split. auto. econstructor. apply star_refl. auto. auto. +- (* final states match *) + intros. destruct s2 as [t s2]; simpl in H1; destruct H1; subst. + inv H. exploit (bsim_match_final_states sim); eauto. eapply star_safe; eauto. + intros [s1'' [A B]]. exists s1''; split; auto. eapply star_trans; eauto. +- (* progress *) + eapply fbs_progress; eauto. +- (* simulation *) + eapply fbs_simulation; eauto. +- (* symbols *) + simpl. exact (bsim_public_preserved sim). +Qed. + +(** Receptiveness of [atomic L]. *) + +Record strongly_receptive {RETVAL: Type} (L: semantics RETVAL) : Prop := + Strongly_receptive { + ssr_receptive: forall s ev1 t1 s1 ev2, + Step L s (ev1 :: t1) s1 -> + match_traces (symbolenv L) (ev1 :: nil) (ev2 :: nil) -> + exists s2, exists t2, Step L s (ev2 :: t2) s2; + ssr_well_behaved: + well_behaved_traces L + }. + +Theorem atomic_receptive: + forall {RETVAL: Type}, + forall (L: _ RETVAL), strongly_receptive L -> receptive (atomic L). +Proof. + intros. constructor; intros. +(* receptive *) + inv H0. + (* silent step *) + inv H1. exists (E0, s'). constructor; auto. + (* start step *) + assert (exists ev2, t2 = ev2 :: nil). inv H1; econstructor; eauto. + destruct H0 as [ev2 EQ]; subst t2. + exploit @ssr_receptive; eauto. intros [s2 [t2 P]]. + exploit @ssr_well_behaved. eauto. eexact P. simpl; intros Q. + exists (t2, s2). constructor; auto. + (* continue step *) + simpl in H2; destruct H2. + assert (t2 = ev :: nil). inv H1; simpl in H0; tauto. + subst t2. exists (t, s0). constructor; auto. simpl; auto. +(* single-event *) + red. intros. inv H0; simpl; omega. +Qed. + +(** * Connections with big-step semantics *) + +(** The general form of a big-step semantics *) + +Record bigstep_semantics (RETVAL: Type) : Type := + Bigstep_semantics { + bigstep_terminates: trace -> RETVAL -> Prop; + bigstep_diverges: traceinf -> Prop + }. + +(** Soundness with respect to a small-step semantics *) + +Record bigstep_sound {RETVAL: Type} (B: bigstep_semantics RETVAL) (L: semantics RETVAL) : Prop := + Bigstep_sound { + bigstep_terminates_sound: + forall t r, + bigstep_terminates B t r -> + exists s1, exists s2, initial_state L s1 /\ Star L s1 t s2 /\ final_state L s2 r; + bigstep_diverges_sound: + forall T, + bigstep_diverges B T -> + exists s1, initial_state L s1 /\ forever (step L) (globalenv L) s1 T +}. + +*) diff --git a/src/backend/Statements/StmtMiniC.v b/src/backend/Statements/StmtMiniC.v new file mode 100644 index 0000000..c7f94b0 --- /dev/null +++ b/src/backend/Statements/StmtMiniC.v @@ -0,0 +1,70 @@ +Require Import cclib.Coqlib. +Require Import cclib.Errors. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.AST. +Require Import backend.Expressions.ExpMiniC. +Require Import backend.Environments.Globalenvs. +Require Import backend.Ctypes. +Require Import backend.Cop. +Require Import backend.MachineModel. + +Inductive statement : Type := + | Sskip : statement (**r do nothing *) + | Sassign : expr -> expr -> statement (**r assignment to evm storage. [lvalue = rvalue], so the left side is computed as pointer. *) + | Sset : ident -> expr -> statement (**r assignment to local env, i.e. the EVM stack. [tempvar = rvalue] *) + (* Clight's Scall does a full internal function call. (as opposed to a method call) + Clike splits this into the simpler tasks of pushing arguments and saving the return value *) + | Scall: option ident (* where the return value goes *) -> label (*name of function to call. *) -> list expr -> statement (**r function call *) + | Ssequence : statement -> statement -> statement (**r sequence *) + | Sifthenelse : expr -> statement -> statement -> statement (**r conditional *) + | Sloop: statement -> statement (**r infinite loop *) + | Sbreak : statement (**r [break] statement *) + | Sreturn : option ident -> statement (**r [return] statement with var for return *) + + (* custom ethereum statements *) + | Stransfer : expr -> expr -> statement (** address.transfer(value) *) + (* call public external function on address, identified by an int, given a value (the amount of money you send), with multiple return values and arguments *) + | Scallmethod : expr (*address*) -> list ident (* return values *) -> int (* method signature, a hash of function name and arguments*) -> expr (* value*) -> option expr (* gas *) -> list expr (* args*) -> statement + | Slog : list expr -> list expr -> statement (** LOG0-6. The first argument is a list of "topics" (at most 6), the second is a list of non-indexed arguments. *) + | Srevert : statement + . + +(** The C loops are derived forms. *) + +Definition Swhile (e: expr) (s: statement) := + Sloop (Sifthenelse e s Sbreak). + +Definition Sdowhile (s: statement) (e: expr) := + Sloop (Ssequence s (Sifthenelse e Sskip Sbreak)). + +(* for (s1; e; s4) s3; *) +Definition Sfor (s1: statement) (e2: expr) (s3: statement) (s4: statement) := + Ssequence s1 (Sloop (Sifthenelse e2 (Ssequence s3 s4) Sbreak)). + +(** ** Functions *) + +(** A function definition is composed of its return type ([fn_return]), + the names and types of its parameters ([fn_params]), the names + and types of its temps ([fn_temps]), and the body of the + function (a statement, [fn_body]). *) + +Record function : Type := mkfunction { + fn_return: type; + fn_params: list (ident * type); + fn_temps: list (ident * type); + fn_locals: list (ident * type); + fn_body: statement +}. + +Definition var_names (vars: list(ident * type)) : list ident := + List.map (@fst ident type) vars. + +(** The semantics uses two environments. The global environment + maps names of functions and global variables to memory block references, + and function pointers to their definitions. (See module [Globalenvs].) *) + +Definition genv := Genv.t function type. + +(* This function is for convenience when writing tests. *) +Definition new_genv := Genv.new_genv function type. diff --git a/src/backend/SymbolicKeccak.v b/src/backend/SymbolicKeccak.v new file mode 100755 index 0000000..3fdd8ee --- /dev/null +++ b/src/backend/SymbolicKeccak.v @@ -0,0 +1,63 @@ +(* Ethereum contracts use Keccak256 hashes a lot, both as an operation +available to the user, and in order to implement hash mappings etc. + +In order for the compiled code to run correctly, we must assume that +there are no hash collissions. This is of course not provable for the +real hash implementation, so instead we use the "symbolic" model +defined in this file, where the hash functions just construct values +in an algebraic data type. Eventually, there will be a compiler phase +that replaces the symbolic hashes with the real ones, withan axiom +stating that this does not introduce collissions. +*) + +Require Import cclib.Integers. +Require Import backend.IndexLib. +Require Import backend.Values.HighValues. + +Definition sha_1 v := + Vhash v. + +Definition sha_2 v1 v2 := + Vhash2 v1 v2. + +Lemma sha_1_range : forall v i, sha_1 v <> Vint i. +Proof. + intros. + unfold sha_1. + congruence. +Qed. + +Lemma sha_2_range : forall v v' i, sha_2 v v' <> Vint i. +Proof. + intros. + unfold sha_2; destruct v; destruct v'; simpl; + congruence. +Qed. + +Lemma sha_1_sha_2_range : forall v v' u, sha_2 v v' <> sha_1 u. +Proof. + intros. + unfold sha_1, sha_2. + destruct v; destruct v'; simpl; congruence. +Qed. + +Lemma sha_1_injective : forall v u, + sha_1 v = sha_1 u -> + v = u. +Proof. + unfold sha_1. + intros; congruence. +Qed. + +Lemma sha_2_injective : forall v1 v2 u1 u2, + sha_2 v1 v2 = sha_2 u1 u2 -> + v1 = u1 /\ v2 = u2. +Proof. + intros. + destruct v1; destruct v2; destruct u1; destruct u2; simpl in * ; split; inversion H; auto. +Qed. + +(* The function must satisfy this property: + me_query_SHA2Pointers: forall a1p a2i, + me_query (Qcall2 Bsha_2 (Vptr a1p) (Vint a2i)) = (Vptr (Ihash a1p a2i)) +*) diff --git a/src/backend/TempModel.v b/src/backend/TempModel.v new file mode 100755 index 0000000..c3efc99 --- /dev/null +++ b/src/backend/TempModel.v @@ -0,0 +1,504 @@ + +(* Model of function temps (and args) *) + +Require Import cclib.Coqlib. +Require Import backend.AST. +Require Import backend.Ctypes. +Require Import cclib.Maps. +Require Import backend.Environments.AllEnvironments. +Require Import backend.Values.HighValues. +Require Import backend.Options. +Require Import cclib.Integers. + +(* All the different intermediate languages (down to stack, which gets rid of temps) represent args and temps as a list of this form: *) +Definition typed_idents : Type := list (positive * type). + +Section FILTER. + + (* The peq in the standard library is opaque, but for us it's better if + it computes. *) + Definition my_peq : forall x y : positive, {x = y} + {x <> y}. + decide equality. + Defined. + +(* filter t by removing ones with duplicate identifiers and removing the ones from excl. + +This is used in the semantics whenever you use temps, in order to +enforce that there are no duplicates. E.g. at the top level, when +estimating gas for calling a function, the compiled code will not push +duplicates, so it also calls this filter function to get a tighter +bound. + +*) +Fixpoint unique_temps_excl (t: typed_idents) (excl: typed_idents) : typed_idents := + match t with nil => nil + | (tmp, ty) :: rest => let utemps := unique_temps_excl rest excl in + match in_dec my_peq tmp (map fst utemps) with + | left tmp_in => utemps (* found in the rest of t, remove duplicate *) + | right tmp_notin => match in_dec my_peq tmp (map fst excl) with + | left tmp_excl => utemps (* found in excl, remove *) + | right tmp_notexcl => (tmp, ty) :: utemps + end + end + end. + +Lemma unique_temps_excl_norepet: forall t excl, + list_norepet (map fst (unique_temps_excl t excl)). +Proof. +induction t; simpl; intros. +constructor. +destruct a. +destruct (in_dec my_peq p (map fst (unique_temps_excl t excl))); auto. +destruct (in_dec my_peq p (map fst excl)); auto. +simpl. constructor; auto. +Qed. + +Lemma unique_temps_excl_in: forall t excl x, + In x (map fst (unique_temps_excl t excl)) -> + In x (map fst t). +Proof. +induction t; simpl; intros. +auto. +destruct a. +destruct (in_dec my_peq p (map fst (unique_temps_excl t excl))). +apply IHt in H. auto. +destruct (in_dec my_peq p (map fst excl)). +apply IHt in H. auto. +simpl in H. +destruct H. subst. auto. +apply IHt in H. auto. +Qed. + +Lemma unique_temps_excl_disjoint: forall t excl, + list_disjoint (map fst (unique_temps_excl t excl)) (map fst excl). +Proof. +induction t; simpl; intros. +unfold list_disjoint. simpl. auto. +destruct a. +destruct (in_dec my_peq p (map fst (unique_temps_excl t excl))). +apply IHt. +destruct (in_dec my_peq p (map fst excl)). +apply IHt. +simpl. +unfold list_disjoint; simpl; intros. +destruct H; intro; subst. contradiction. +unfold list_disjoint in IHt. assert (A := IHt excl y y H H0). auto. +Qed. + +Lemma unique_temps_excl_defined : forall i ty lst excl, + In (i,ty) lst -> + ~In i (map fst excl) -> + In i (map fst (unique_temps_excl lst excl)). +Proof. + induction lst. + - intros. + simpl in *. + auto. + - intros. + simpl in *. + destruct a. + destruct (in_dec my_peq p (map fst (unique_temps_excl lst excl))). + + destruct H. + * inversion H. + subst. + assumption. + * apply IHlst; + assumption. + + destruct H. + * inversion H. + subst. + destruct (in_dec my_peq i (map fst excl)). + tauto. + simpl; auto. + * destruct (in_dec my_peq p (map fst excl)). + apply IHlst; assumption. + simpl. right. apply IHlst; assumption. +Qed. + +Definition unique_temps (t: typed_idents) : typed_idents := + unique_temps_excl t nil. + +Lemma unique_temps_disjoint: forall t excl, + list_disjoint (map fst (unique_temps_excl t excl)) (map fst (unique_temps excl)). +Proof. +unfold list_disjoint. intros. intro; subst. +apply unique_temps_excl_in in H0. +assert (A := unique_temps_excl_disjoint t excl y y H H0). auto. +Qed. + +Definition unique_all_temps (t: typed_idents) (t': typed_idents) : typed_idents := + unique_temps_excl t t' ++ unique_temps t'. + +Lemma norepet_disjoint: forall A (t t': list A), + list_disjoint t t' -> + list_norepet t -> + list_norepet t' -> + list_norepet (t ++ t'). +Proof. +induction t; simpl; intros. +auto. +inv H0. +assert (~ In a t'). { intro. unfold list_disjoint in H. unfold not in H. eapply H; eauto. simpl. auto. } +apply list_disjoint_cons_left in H. +constructor. +intro. apply in_app_or in H2. destruct H2; contradiction. +apply IHt; auto. +Qed. + +Lemma unique_all_temps_norepet: forall t t', + list_norepet (map fst (unique_all_temps t t')). +Proof. +unfold unique_all_temps. +intros. +rewrite map_app. +apply norepet_disjoint. +apply unique_temps_disjoint. +apply unique_temps_excl_norepet. +apply unique_temps_excl_norepet. +Qed. + +End FILTER. + +Section INIT. + +(* Functions and lemmas related to setting the initial args in the temp_env. *) + +(* +init_args is the first part of the semantics of calling a function, then init_temps adds zeros for actual temporaries. + +For proof's sake, add these to the temp_env in the order that they get pushed to the stack. +Outputs None if args and vals are different lengths. + *) +Fixpoint init_args (args: typed_idents) (vals: list val) : option temp_env := + match args, vals with + | nil, nil => Some (PTree.empty _) + | (name, _)::rest, v::vs => r <- init_args rest vs ;; Some ((PTree.set name v r):temp_env) + | _, _ => None + end. + +Remark init_args_len: forall args vals le, + init_args args vals = Some le -> length args = length vals. +Proof. induction args; induction vals; simpl; intros. +auto. discriminate. destruct a; discriminate. destruct a. +BindSome H r r_eq. SomeSome H le. replace (length vals) with (length args). +auto. eapply IHargs; eauto. +Qed. + +Lemma init_args_defined : forall (args: typed_idents) (vals: list val) i ty t, + In (i, ty) args -> + init_args args vals = Some t -> + exists v, PTree.get i t = Some v. +Proof. + induction args; destruct vals; simpl; try congruence; try tauto. + - destruct a; intros; simpl in *. + congruence. + - intros. + destruct H. + + subst; simpl in *. + destruct (init_args args vals) as [t'|]; simpl in *; [|congruence]. + simpl in *. + inversion H0. + subst. + exists v. + rewrite PTree.gss. + reflexivity. + + destruct a as [i' v']. + destruct (init_args args vals) as [t'|] eqn:H_eqt'; simpl in *; [|congruence]. + specialize (IHargs vals _ ty t' H H_eqt'). + inversion H0. + destruct (peq i' i). + * subst. + rewrite PTree.gss. + eauto. + * rewrite PTree.gso by congruence. + exact IHargs. +Qed. + +Fixpoint init_temps (temps: typed_idents) (base: temp_env) : temp_env := + match temps with + | nil => base + | (name, _)::rest => init_temps rest (PTree.set name (Vint Int256.zero) base) + end. + + +End INIT. + +(* generic function with temps *) +Record ftemps : Type := mkftemps { + f_t : Type; + f_temps : f_t -> typed_idents; + f_args : f_t -> typed_idents +}. + +Section FUNCTIONTYPE. + +(* All the intermediate langauges have a notion of arguments and temps, so here we prove some lemmas about those. *) + + +Variable ftype : ftemps. +Definition function : Type := f_t ftype. +Definition fn_temps := f_temps ftype. +Definition fn_params := f_args ftype. + +Section FUNCTION. + +Variable f : function. +Definition temps := fn_temps f. +Definition params := fn_params f. + +(* The function temps and args after removing duplicates. *) +Definition some_temps : typed_idents := + unique_temps_excl temps params. + +Definition some_args : typed_idents := + unique_temps params. + + +Lemma some_args_defined : forall i ty, + In (i, ty) params -> + exists ty', In (i,ty') some_args. +Proof. + unfold some_args. + intros i ty. + induction params. + - simpl; intros; tauto. + - simpl. + unfold unique_temps. + simpl. + intros [H | H]. + + subst. + destruct (in_dec my_peq i (map fst (unique_temps_excl t nil))) as [i0 | i0]. + * apply list_in_map_inv in i0. + destruct i0 as [[i' ty'] [Hi' H']]. + simpl in Hi'. + subst. + exists ty'. + exact H'. + * exists ty. + simpl. + auto. + + specialize (IHt H). + destruct a as [i0 ty0]. + assert (Hin: In i (map fst (unique_temps_excl t nil))). + { + apply unique_temps_excl_defined with ty. + assumption. + simpl; auto. + } + apply list_in_map_inv in Hin. + destruct Hin as [[i' ty'] [Hin1 Hin2]]. + destruct (in_dec my_peq i0 (map fst (unique_temps_excl t nil))). + * simpl in *; subst. + exists ty'. + assumption. + * simpl in *; subst. + exists ty'. + right. + assumption. + Qed. + + (* +When combining temps, we want a few properties which are required by how they are used on the Stack. +1. The temps should be have unique identifiers. This allows random access sets and gets, with nice proofs. +2. Args are pushed first, and temps pushed after, so in list form they are (temps ++ args) +*) +Definition all_temps : typed_idents := + unique_all_temps temps params. + +Remark all_temps_split: all_temps = some_temps ++ some_args. +Proof. auto. Qed. + +(* Returns None if the length of the argument list is different form the number of values provided. *) +Definition function_initial_temps (vals: list val) : option temp_env := + args_initial <- init_args some_args vals ;; + Some (init_temps some_temps args_initial). + +Lemma init_temps_defined : forall i ty lst te, + In (i, ty) lst \/ PTree.get i te <> None -> + exists v, PTree.get i (init_temps lst te) = Some v. +Proof. + induction lst. + - inversion 1. + + inversion H0. + + simpl. + destruct (PTree.get i te) as [|v]. + * eexists; reflexivity. + * congruence. + - intros. + simpl in H. + destruct H as [[H|H] | H]. + + subst. + simpl. + apply IHlst. + right. + rewrite PTree.gss. + congruence. + + destruct a as [a aty]. + simpl. + destruct (IHlst (PTree.set a (Vint Int256.zero) te)) as [v Hv]. + left; assumption. + exists v; apply Hv. + + destruct a as [a aty]. + simpl. + destruct (IHlst (PTree.set a (Vint Int256.zero) te)) as [v Hv]. + { right. + destruct (my_peq a i). + - subst. + rewrite PTree.gss. + congruence. + - rewrite PTree.gso; congruence. + } + exists v; apply Hv. +Qed. + + + + +End FUNCTION. + +End FUNCTIONTYPE. + + +(* match two functions that have the same temps but different type. + + This is used in the proofs when going between two intermediate + languages which both have functions with temps and args. The lemmas + says that the corresponding enviroments in the two langauges are the + same. + *) +Section FUNCTIONMATCH. + +Variable ftype1 : ftemps. +Definition function1 : Type := f_t ftype1. +Definition fn_temps1 := f_temps ftype1. +Definition fn_params1 := f_args ftype1. + +Variable ftype2 : ftemps. +Definition function2 : Type := f_t ftype2. +Definition fn_temps2 := f_temps ftype2. +Definition fn_params2 := f_args ftype2. + + +Inductive match_ftemps: function1 -> function2 -> Prop := + | match_ftemps_intro: forall f f', + fn_temps1 f = fn_temps2 f' -> + fn_params1 f = fn_params2 f' -> + match_ftemps f f'. + +Lemma match_some_temps: forall f f', + match_ftemps f f' -> + some_temps ftype1 f = some_temps ftype2 f'. +Proof. +intros. inv H. unfold some_temps. +replace (temps ftype1 f) with (temps ftype2 f'). +replace (params ftype1 f) with (params ftype2 f'). +auto. +Qed. + +Lemma match_some_args: forall f f', + match_ftemps f f' -> + some_args ftype1 f = some_args ftype2 f'. +Proof. +intros. inv H. unfold some_args. +replace (params ftype1 f) with (params ftype2 f'). +auto. +Qed. + +Lemma match_all_temps: forall f f', + match_ftemps f f' -> + all_temps ftype1 f = all_temps ftype2 f'. +Proof. +intros. inv H. unfold all_temps. +replace (temps ftype1 f) with (temps ftype2 f'). +replace (params ftype1 f) with (params ftype2 f'). +auto. +Qed. + +Lemma match_initial_temps: forall f f' vals le, + match_ftemps f f' -> + function_initial_temps ftype1 f vals = Some le -> + function_initial_temps ftype2 f' vals = Some le. +Proof. +unfold function_initial_temps. +intros. BindSome H0 args_initial args_eq. SomeSome H0 le. +replace (some_args ftype2 f') with (some_args ftype1 f). +replace (some_temps ftype2 f') with (some_temps ftype1 f). +rewrite args_eq. simpl. auto. +apply match_some_temps; auto. apply match_some_args; auto. +Qed. + + +End FUNCTIONMATCH. + +Lemma function_initial_temps_defined : forall i ty (ftype : ftemps) (f :function ftype) t, + In (i, ty) (fn_temps ftype f) -> + ~In i (map fst (params ftype f)) -> + exists v, PTree.get i + (init_temps (some_temps ftype f) + t) + = Some v. +Proof. + intros i ty ftype f t Hin Hnin. + assert (in_some_temps + := unique_temps_excl_defined i ty (fn_temps ftype f) _ Hin Hnin). + apply list_in_map_inv in in_some_temps. + destruct in_some_temps as [ [i' ty'] [i'_eq Hin']]. + subst. + + destruct (init_temps_defined i' ty' (some_temps ftype f) t) as [v H]. + { left. + exact Hin'. + } + exists v. + exact H. +Qed. + + +Definition function_initial_temps_defined2 + : forall (ftype:ftemps) (f:function ftype) + (vals: list val) (i:positive) (ty:type) t, + In (i, ty) (params ftype f) -> + function_initial_temps ftype f vals = Some t -> + exists v, PTree.get i t = Some v. +Proof. + intros ftype f vals i ty t Hin Hinitial_temps. + unfold function_initial_temps in Hinitial_temps. + destruct (init_args (some_args ftype f) vals) as [t'|] eqn:Hinit_args; simpl in *; [|congruence] . + assert (Hin': exists ty', In (i,ty') (some_args ftype f) ). + { + apply some_args_defined with ty. + exact Hin. + } + destruct Hin' as [ty' Hin']. + destruct (init_args_defined (some_args ftype f) vals i ty' t' Hin' Hinit_args). + destruct (init_temps_defined i ty' (some_temps ftype f) t') as [v Hv]. + { right. + congruence. + } + inversion Hinitial_temps. + exists v. + exact Hv. +Qed. + + +(* Helper function for (possibly) setting temps.*) +Definition optset (key: option ident) (v: val) (le: temp_env) : temp_env := + match key with None => le + | Some id => PTree.set id v le + end. + +(* Helper function to set a list of keys to a list of values. +For the sake of the proof, order matters. The first ident/value pair is set first. +Not a total function because the lists must be the same length *) +Fixpoint listset (keys: list ident) (vals: list val) (le: temp_env) : option temp_env := + match keys, vals with + | nil, nil => Some le + | k::ks, v::vs => listset ks vs (PTree.set k v le) + | _, _ => None + end. + +Remark listset_len: forall rv_keys rvs le le', + listset rv_keys rvs le = Some le' -> length rvs = length rv_keys. +Proof. induction rv_keys; induction rvs; intros; inv H. auto. simpl. +replace (length rv_keys) with (length rvs). auto. eapply IHrv_keys; eauto. +Qed. diff --git a/src/backend/Values/HighValues.v b/src/backend/Values/HighValues.v new file mode 100755 index 0000000..50cc9ea --- /dev/null +++ b/src/backend/Values/HighValues.v @@ -0,0 +1,258 @@ +(* This may need to be changed quite a bit, see the readme file. + +*) + +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** This module defines the type of values that is used in the dynamic + semantics of all our intermediate languages. *) + +Require Import cclib.Coqlib. +Require Import cclib.Integers. +Require Import cclib.Maps. +Require Import backend.AST. +Require Import backend.IndexLib. +Require Import backend.Values.LowValues. + + +(* we want maps that are initialized to 0 by default, +which requires an IMap instead of a PTree *) +Module IdentIndexed <: INDEXED_TYPE. + +Definition t := positive. +Definition index (i: positive) : positive := i. +Definition index_inv (i: positive) : positive := i. + +Lemma index_invertible : forall x, index_inv (index x) = x. +Proof. reflexivity. Qed. + +Lemma index_inj: forall (x y: positive), index x = index y -> x = y. +Proof. +simpl. intros x y H. exact H. +Qed. +Definition eq := peq. + +End IdentIndexed. + +Module IdentMap := IMap(IdentIndexed). + +(* Extended identifier, +so a single identifier (i.e. array name) +can refer to multiple values *) + +Section ExtendedIdent. + + (*Extended Ident 2: + An extended identifier has a base that is either Global or Local, + with a positive integer used to distinguish local contexts. From + the base new extended identifiers are generated by indexing into + an array or accessing an identifier as a field in a struct (treating + top-level variables as fields as well.) *) + + Inductive ident_ext : Type := + | Global: ident_ext + | Local: int256 -> ident_ext + | Field: ident_ext -> ident -> ident_ext + | Index: ident_ext -> int256 -> ident_ext + . + +End ExtendedIdent. + +(* Allow extended identifiers to be used as the index of a map. *) +Module IdentExtIndexed <: INDEXED_TYPE. + + Definition t := ident_ext. + + Fixpoint index (eid : ident_ext) : positive := + match eid with + | Global => xH + | Local ctx => int_index ctx + | Field eid id => inject_positive (index eid) id + | Index eid i => inject_positive (index eid) (int_index i) + end. + +Definition index_inv (p : positive) : ident_ext. +Admitted. + +Lemma index_invertible : forall x, index_inv (index x) = x. +Admitted. + +Lemma index_inj: forall (x y: ident_ext), index x = index y -> x = y. +Proof. +Admitted. +(* induction x; intro y; destruct y; try discriminate. +- + unfold index. intro H; injection H; clear H. + intro H; rewrite H; reflexivity. +- + simpl. intro H; injection H; clear H. + intro. + assert ((index x)=(index y)). + apply injective_positive_1 with (int_index i) (int_index i0). exact H. + replace y with x. + assert ((int_index i)=(int_index i0)). + apply injective_positive_2 with (index x) (index y). exact H. clear H. + replace i0 with i. + reflexivity. + apply int_index_injective; assumption. + apply IHx; assumption. +Qed.*) + +Lemma eq: forall (x y: ident_ext), {x = y} + {x <> y}. +Proof. + induction x; destruct y; intros; decide equality; + try decide equality; try apply Int256.eq_dec; + try apply Pos.eq_dec. +Qed. + +End IdentExtIndexed. + +Module IdentExtMap := IMap(IdentExtIndexed). + +(* `ident_ext_extends i1 i2` is true if `i2` is of the form + `Ihash (Ihash ... i1 ... ofs') ofs`. In other words, i2 + designates a sub-part of the object located at i1. *) +Inductive ident_ext_extends : forall (i1 i2 : ident_ext), Prop := +| IIE_refl : forall eid, ident_ext_extends eid eid +| IIE_index : forall eid1 eid2 i, + ident_ext_extends eid1 eid2 -> + ident_ext_extends eid1 (Index eid2 i) +| IIE_field : forall eid1 eid2 id, + ident_ext_extends eid1 eid2 -> + ident_ext_extends eid1 (Field eid2 id). + +Definition ident_ext_extends_inversion': forall i i2, + ident_ext_extends i i2 -> + forall i1 o f, (i = (Index i1 o) \/ i = (Field i1 f)) -> + ident_ext_extends i1 i2. +Proof. + induction 1; destruct 1; subst; constructor. + - constructor. + - constructor. + - unshelve (eapply IHident_ext_extends; eauto); (exact (1%positive) || exact Int256.zero). + - unshelve (eapply IHident_ext_extends; eauto); (exact (1%positive) || exact Int256.zero). + - unshelve (eapply IHident_ext_extends; eauto); (exact (1%positive) || exact Int256.zero). + - unshelve (eapply IHident_ext_extends; eauto); (exact (1%positive) || exact Int256.zero). +Qed. + +Lemma ident_ext_extends_inversion_Index: forall i1 i2 o, + ident_ext_extends (Index i1 o) i2 -> + ident_ext_extends i1 i2. +Proof. + intros; eapply ident_ext_extends_inversion' with (f:= 1%positive); eauto. +Qed. + +Lemma ident_ext_extends_inversion_Field: forall i1 i2 f, + ident_ext_extends (Field i1 f) i2 -> + ident_ext_extends i1 i2. +Proof. + intros; eapply ident_ext_extends_inversion' with (o := Int256.zero); eauto. +Qed. + +Fixpoint ident_ext_length i := + match i with + | Global => O + | Local _ => O + | Field i _ => S (ident_ext_length i) + | Index i _ => S (ident_ext_length i) + end. + +Lemma ident_ext_extends_longer : forall i j, + ident_ext_extends i j -> + (ident_ext_length i <= ident_ext_length j)%nat. +Proof. + induction 1; simpl; omega. +Qed. + +Lemma ident_ext_extends_disjoint_Index : forall o1 o2, + o1 <> o2 -> + forall i' i0, + ident_ext_extends (Index i0 o1) i' -> + ~ ident_ext_extends (Index i0 o2) i'. +Proof. + intros o1 o2 Hneq. + unfold not. + induction i'; intros i0 H1 H2. + - inversion H1. + - inversion H1. + - inversion H1; inversion H2; subst. + apply IHi' with i0; auto. + - inversion H1; inversion H2; subst. + + congruence. + + apply ident_ext_extends_longer in H6. + simpl in H6. + omega. + + apply ident_ext_extends_longer in H3. + simpl in H3. + omega. + + unfold not in IHi'; eapply IHi'; eassumption. +Qed. + +Fixpoint ident_ext_base (eid : ident_ext) : ident := + match eid with + | Field Global id => id + | Field (Local _) id => id + | Field eid _ => ident_ext_base eid + | Index eid _ => ident_ext_base eid + | Global => 1%positive (* dummy, should not happen *) + | Local _ => 1%positive (* dummy, should not happen *) + end. + + + +Section HValue. + + (** A high value extends low values with extended-identifier pointers: *) + Inductive val : Type := + | Vunit: val + | Vint: int256 -> val + | Vptr: lval -> val + | Vhash: val -> val + | Vhash2: val -> val -> val + with lval := + | LVeid: ident_ext -> lval + | LVhash: int256 -> lval + | LVhash2: lval -> int256 -> lval. + + Definition Vzero: val := Vint Int256.zero. + Definition Vone: val := Vint Int256.one. + Definition Vmone: val := Vint Int256.mone. + + Definition Vtrue: val := Vint Int256.one. + Definition Vfalse: val := Vint Int256.zero. + + Fixpoint OfLow (v:LowValues.val) : val := + match v with + | LowValues.Vunit => Vunit + | LowValues.Vint i => Vint i + | LowValues.Vhash v => Vhash (OfLow v) + | LowValues.Vhash2 v1 v2 => Vhash2 (OfLow v1) (OfLow v2) + end. + + Fixpoint ToLowErr (v:val) : option LowValues.val := + match v with + | Vunit => Some LowValues.Vunit + | Vint i => Some (LowValues.Vint i) + | Vhash v => option_map LowValues.Vhash (ToLowErr v) + | Vhash2 v1 v2 => + match ToLowErr v1, ToLowErr v2 with + | Some v1', Some v2' => Some (LowValues.Vhash2 v1' v2') + | _,_ => None + end + | Vptr _ => None + end. + +End HValue. diff --git a/src/backend/Values/LowValues.v b/src/backend/Values/LowValues.v new file mode 100644 index 0000000..60f9a7d --- /dev/null +++ b/src/backend/Values/LowValues.v @@ -0,0 +1,140 @@ +(* This may need to be changed quite a bit, see the readme file. + +*) + +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** This module defines the type of values that is used in the dynamic + semantics of all our intermediate languages. *) + +Require Import cclib.Coqlib. +Require Import cclib.Integers. +Require Import cclib.Maps. +Require Import backend.AST. +Require Import backend.IndexLib. + + +(* we want maps that are initialized to 0 by default, +which requires an IMap instead of a PTree *) +Module IdentIndexed <: INDEXED_TYPE. + +Definition t := positive. +Definition index (i: positive) : positive := i. +Definition index_inv (i: positive) : positive := i. + +Lemma index_invertible : forall x, index_inv (index x) = x. +Proof. reflexivity. Qed. + +Lemma index_inj: forall (x y: positive), index x = index y -> x = y. +Proof. +simpl. intros x y H. exact H. +Qed. +Definition eq := peq. + +End IdentIndexed. + +Module Int256Map := IMap(Int256Indexed). + +(* Currently skipping hashes; should come back to do the rest. *) +(* `ident_ext_extends i1 i2` is true if `i2` is of the form + `Ihash (Ihash ... i1 ... ofs') ofs`. In other words, i2 + designates a sub-part of the object located at i1. *) +Inductive ident_ext_extends : int256 -> int256 -> Prop := +| IIE_refl : forall i, ident_ext_extends i i. + +(*| IIE_index : forall eid1 eid2 i, + ident_ext_extends eid1 eid2 -> + ident_ext_extends eid1 (Index eid2 i) +| IIE_field : forall eid1 eid2 id, + ident_ext_extends eid1 eid2 -> + ident_ext_extends eid1 (Field eid2 id).*) + +(*Definition ident_ext_extends_inversion': forall i i2, + ident_ext_extends i i2 -> + forall i1 o, i = (Ihash i1 o) -> + ident_ext_extends i1 i2. +Proof. + induction 1; intros; subst. + - constructor; constructor. + - constructor. + eapply IHident_ext_extends. + reflexivity. +Qed. + +Lemma ident_ext_extends_inversion: forall i1 i2 o, + ident_ext_extends (Ihash i1 o) i2 -> + ident_ext_extends i1 i2. +Proof. + intros; eauto using ident_ext_extends_inversion'. +Qed. + +Fixpoint ident_ext_length i := + match i with + | Iident _ => O + | Ihash i _ => S (ident_ext_length i) + end. + +Lemma ident_ext_extends_longer : forall i j, + ident_ext_extends i j -> + (ident_ext_length i <= ident_ext_length j)%nat. +Proof. + induction 1. + - omega. + - simpl. + omega. + Qed. + +Lemma ident_ext_extends_disjoint : forall o1 o2, + o1 <> o2 -> + forall i' i0, + ident_ext_extends (Ihash i0 o1) i' -> + ~ ident_ext_extends (Ihash i0 o2) i'. +Proof. + intros o1 o2 Hneq. + unfold not. + induction i'; intros i0 H1 H2. + - inversion H1. + - inversion H1; inversion H2; subst. + + congruence. + + apply ident_ext_extends_longer in H6. + simpl in H6. + omega. + + apply ident_ext_extends_longer in H3. + simpl in H3. + omega. + + unfold not in IHi'; eapply IHi'; eassumption. +Qed. + *) + +(** A value is either: +- Vunit. When a function returns unit, we compile it to code which actually pushes a value to the stack, namely this value. +- a machine integer; +- a pointer, which is also a machine integer; +- a hash + *) +Inductive val: Type := +| Vunit: val +| Vint: int256 -> val +| Vhash: val -> val +| Vhash2: val -> val -> val +. + +Definition Vzero: val := Vint Int256.zero. +Definition Vone: val := Vint Int256.one. +Definition Vmone: val := Vint Int256.mone. + +Definition Vtrue: val := Vint Int256.one. +Definition Vfalse: val := Vint Int256.zero. diff --git a/src/backend/extraction/ASM.ml b/src/backend/extraction/ASM.ml index 00283b7..a434868 100755 --- a/src/backend/extraction/ASM.ml +++ b/src/backend/extraction/ASM.ml @@ -183,7 +183,7 @@ let getlabel id s = "\"" ^ (spos id) ^ " : done \"" | Stransfer (e1, e2, n, n') -> "\"" ^ (spos id) ^ " : transfer \"" - | StmtCGraph.Scallmethod (eil, il, i, e, el, n, n') -> + | StmtCGraph.Scallmethod (eil, il, i, e, eo, el, n, n') -> "\"" ^ (spos id) ^ " : callmethod \"" | Slog (el1, el2, n) -> "\"" ^ (spos id) ^ " : log \"" @@ -229,7 +229,7 @@ let mnemonics_cgraph cd en = label ^ "\n" | Stransfer (e1, e2, n1, n2) -> label ^ "\n" ^ transitionsn n1 ^ "\n" ^ transitionsn n2 - | Scallmethod (eil, il, i, e, el, n1, n2) -> + | Scallmethod (eil, il, i, e, eo, el, n1, n2) -> label ^ "\n" ^ transitionsn n1 ^ "\n" ^ transitionsn n2 | Slog (el1, el2, n) -> label ^ "\n" ^ transitionsn n @@ -386,7 +386,7 @@ let assemble_inst programsize = function | EVM_JUMPDEST l -> "5b" | EVM_PUSH (n, data) -> (assemble_op (95 + n)) ^ data - | EVM_DUP n -> assemble_op (127 + n) + | EVM_DUP n -> (* print_endline (string_of_int n); if n > 16 then "DUP error" else *) assemble_op (127 + n) | EVM_SWAP n -> assemble_op (143 + n) | EVM_LOG n -> assemble_op (160 + n) | EVM_CALL -> "f1" diff --git a/src/backend/extraction/BytecodeExt.ml b/src/backend/extraction/BytecodeExt.ml index 75a732f..f071e30 100644 --- a/src/backend/extraction/BytecodeExt.ml +++ b/src/backend/extraction/BytecodeExt.ml @@ -26,9 +26,11 @@ let bytecode runtime genv = | false -> assemble asm | true -> assemble asm_runtime -let assembly genv = - let asm , _, _ = get_bytecode_params genv in - mnemonics asm +let assembly runtime genv = + let asm , asm_runtime, _ = get_bytecode_params genv in + match runtime with + | false -> mnemonics asm + | true -> mnemonics asm_runtime let ewasm runtime genv = match Glue.full_compile_genv_wasm genv with diff --git a/src/backend/extraction/CopExt.ml b/src/backend/extraction/CopExt.ml index 58a6d7f..d3160b1 100755 --- a/src/backend/extraction/CopExt.ml +++ b/src/backend/extraction/CopExt.ml @@ -37,8 +37,8 @@ let show_binop_symb = function | Cop.Odiv -> "/" | Cop.Omod -> "%" | Cop.Oexp -> "**" - | Cop.Oand -> "&&" - | Cop.Oor -> "||" + | Cop.Oand -> "&" + | Cop.Oor -> "|" | Cop.Oxor -> "^" | Cop.Oshl -> "<<" | Cop.Oshr -> ">>" diff --git a/src/backend/extraction/Gen.ml b/src/backend/extraction/Gen.ml index 660011b..aa43706 100644 --- a/src/backend/extraction/Gen.ml +++ b/src/backend/extraction/Gen.ml @@ -516,6 +516,7 @@ and clike_lvalue = function | Eglob (id, ty) -> ret (Obj.magic coq_Monad_optErr) (Eunop (Osha_1, (constofpos id), (Tpointer (Coq_stor, ty)))) +| Etempvar (id, ty) -> ret (Obj.magic coq_Monad_optErr) (Etempvar (id, ty)) | Ederef (ex0, _) -> clike_rvalue ex0 | Efield (ex0, id, ty) -> bind (Obj.magic coq_Monad_optErr) (clike_lvalue ex0) (fun ex' -> @@ -1147,6 +1148,16 @@ let rec clike_rvalue_list = function bind (Obj.magic coq_Monad_optErr) (clike_rvalue_list tl) (fun rest -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (first, rest)))) +(** val clike_rvalue_option : expr option -> expr option optErr **) + +let rec clike_rvalue_option = function +| Some e -> + let e' = clike_rvalue e in + (match e' with + | Success s -> ret (Obj.magic coq_Monad_optErr) (Some s) + | Error msg -> Error msg) +| None -> ret (Obj.magic coq_Monad_optErr) None + (** val clike_stm : statement -> statement optErr **) let rec clike_stm = function @@ -1179,15 +1190,17 @@ let rec clike_stm = function (fun addr0 -> bind (Obj.magic coq_Monad_optErr) (Obj.magic clike_rvalue val0) (fun val1 -> ret (Obj.magic coq_Monad_optErr) (Stransfer (addr0, val1)))) -| Scallmethod (addr, retvals, funsig, val0, args) -> +| Scallmethod (addr, retvals, funsig, val0, gas, args) -> bind (Obj.magic coq_Monad_optErr) (Obj.magic clike_rvalue addr) (fun addr0 -> bind (Obj.magic coq_Monad_optErr) (Obj.magic clike_rvalue val0) (fun val1 -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic clike_rvalue_list args) - (fun args0 -> - ret (Obj.magic coq_Monad_optErr) (Scallmethod (addr0, retvals, - funsig, val1, args0))))) + bind (Obj.magic coq_Monad_optErr) (Obj.magic clike_rvalue_option gas) + (fun gas0 -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic clike_rvalue_list args) + (fun args0 -> + ret (Obj.magic coq_Monad_optErr) (Scallmethod (addr0, retvals, + funsig, val1, gas0, args0)))))) | Slog (topics, args) -> bind (Obj.magic coq_Monad_optErr) (Obj.magic clike_rvalue_list topics) (fun topics0 -> diff --git a/src/backend/extraction/Gen.mli b/src/backend/extraction/Gen.mli index 83a87c4..874639c 100644 --- a/src/backend/extraction/Gen.mli +++ b/src/backend/extraction/Gen.mli @@ -21,6 +21,8 @@ val clike_lvalue : expr -> expr optErr val clike_rvalue_list : expr list -> expr list optErr +val clike_rvalue_option : expr option -> expr option optErr + val clike_stm : statement -> statement optErr val clike_function : coq_function -> coq_function optErr diff --git a/src/backend/extraction/Gen0.ml b/src/backend/extraction/Gen0.ml index 0677fab..e0ed31f 100644 --- a/src/backend/extraction/Gen0.ml +++ b/src/backend/extraction/Gen0.ml @@ -5131,7 +5131,7 @@ let rec clocal_stm scrmap = function clocal_rvalue scrmap (Obj.magic cont') val0 scr in clocal_rvalue scrmap cont addr O -| Scallmethod (addr, retvals, funsig, val0, args) -> +| Scallmethod (addr, retvals, funsig, val0, gas, args) -> let cont = fun all_exps scr -> match all_exps with | Coq_nil -> @@ -5228,10 +5228,161 @@ let rec clocal_stm scrmap = function EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))) | Coq_cons (_, _) -> ret coq_Monad_optErr (Coq_pair ((StmtClocal.Scallmethod (addr, - retvals, funsig, val0, args)), scr))) + retvals, funsig, val0, gas, args)), scr))) in - clocal_expr_list scrmap (Obj.magic cont) (Coq_cons (addr, (Coq_cons (val0, - args)))) O + let cont' = fun all_exps scr -> + match all_exps with + | Coq_nil -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))) + | Coq_cons (_, l) -> + (match l with + | Coq_nil -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))) + | Coq_cons (_, l0) -> + (match l0 with + | Coq_nil -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))) + | Coq_cons (_, _) -> + ret coq_Monad_optErr (Coq_pair ((StmtClocal.Scallmethod (addr, + retvals, funsig, val0, gas, args)), scr)))) + in + (match gas with + | Some g -> + clocal_expr_list scrmap (Obj.magic cont') (Coq_cons (addr, (Coq_cons + (val0, (Coq_cons (g, args)))))) O + | None -> + clocal_expr_list scrmap (Obj.magic cont) (Coq_cons (addr, (Coq_cons + (val0, args)))) O) | Slog (topics, args) -> let cont = fun all_exps scr -> bind coq_Monad_optErr (Obj.magic split_exps (length topics) all_exps) diff --git a/src/backend/extraction/Gen1.ml b/src/backend/extraction/Gen1.ml index 5d1afb0..cb39c80 100644 --- a/src/backend/extraction/Gen1.ml +++ b/src/backend/extraction/Gen1.ml @@ -1,18 +1,13 @@ -open AST open Ascii open BinNums open BinPos open Coqlib open Ctypes open Datatypes -open ExpCintptr open Globalenvs -open List0 open Maps0 open Monad -open Nat0 open OptErrMonad -open PeanoNat open Specif open StmtCGraph open StmtCintptr @@ -120,1877 +115,11 @@ let rec cgraph_statement s nd nret nrev nbrk = | Sreturn retvar -> add_instr (StmtCGraph.Sreturn (retvar, nret)) | Shash (ex1, ex2, exo) -> add_instr (StmtCGraph.Shash (ex1, ex2, exo, nd)) | Stransfer (a, v) -> add_instr (StmtCGraph.Stransfer (a, v, nrev, nd)) - | Scallmethod (a, rvs, sig0, v, args) -> - add_instr (StmtCGraph.Scallmethod (a, rvs, sig0, v, args, nrev, nd)) + | Scallmethod (a, rvs, sig0, v, g, args) -> + add_instr (StmtCGraph.Scallmethod (a, rvs, sig0, v, g, args, nrev, nd)) | Slog (topics, args) -> add_instr (StmtCGraph.Slog (topics, args, nd)) | Srevert -> ret (Obj.magic coq_Monad_mon) nrev -type set = bool PTree.t - -type coq_LVDomain = set - -(** val set_empty : set **) - -let set_empty = - PTree.empty - -(** val set_union : positive list -> set -> set -> set -> set **) - -let rec set_union keys x y r = - match keys with - | Coq_nil -> r - | Coq_cons (hd, rest) -> - set_union rest x y - (PTree.set hd - (match PTree.get_default Coq_false hd x with - | Coq_true -> Coq_true - | Coq_false -> PTree.get_default Coq_false hd y) r) - -(** val set_minus : positive list -> set -> set -> set -> set **) - -let rec set_minus keys x y r = - match keys with - | Coq_nil -> r - | Coq_cons (hd, rest) -> - set_minus rest x y - (PTree.set hd - (match PTree.get_default Coq_false hd x with - | Coq_true -> negb (PTree.get_default Coq_false hd y) - | Coq_false -> Coq_false) r) - -(** val set_add : ident -> set -> set **) - -let rec set_add i r = - PTree.set i Coq_true r - -(** val set_in : ident -> set -> bool **) - -let rec set_in i r = - PTree.get_default Coq_false i r - -(** val set_card : set -> nat **) - -let rec set_card x = - PTree.fold1 (fun acc e -> - match e with - | Coq_true -> add acc (S O) - | Coq_false -> acc) x O - -(** val set_eq : set -> set -> bool **) - -let rec set_eq x y = - match Nat.eqb (set_card (set_minus (PTree.xkeys x Coq_xH) x y PTree.empty)) - O with - | Coq_true -> - Nat.eqb (set_card (set_minus (PTree.xkeys y Coq_xH) y x PTree.empty)) O - | Coq_false -> Coq_false - -(** val set_fold_left : ('a1 -> ident -> 'a1) -> set -> 'a1 -> 'a1 **) - -let set_fold_left f m v = - PTree.fold (fun acc i e -> - match e with - | Coq_true -> f acc i - | Coq_false -> acc) m v - -(** val is_exit : node -> node -> node -> bool **) - -let is_exit nret nrev n = - match Pos.eqb n nrev with - | Coq_true -> Coq_true - | Coq_false -> Pos.eqb n nret - -(** val get_use_expr : expr -> coq_LVDomain optErr **) - -let rec get_use_expr = function -| Etempvar (i, _) -> ret (Obj.magic coq_Monad_optErr) (set_add i set_empty) -| Esload (e0, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr e0) (fun e' -> - ret (Obj.magic coq_Monad_optErr) e') -| Emload (e0, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr e0) (fun e' -> - ret (Obj.magic coq_Monad_optErr) e') -| Eaddr (e0, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr e0) (fun e' -> - ret (Obj.magic coq_Monad_optErr) e') -| Eunop (_, e0, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr e0) (fun e' -> - ret (Obj.magic coq_Monad_optErr) e') -| Ebinop (_, e1, e2, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr e1) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr e2) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (set_union (app (PTree.xkeys e1' Coq_xH) (PTree.xkeys e2' Coq_xH)) - e1' e2' PTree.empty))) -| Ecall1 (_, e0, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr e0) (fun e' -> - ret (Obj.magic coq_Monad_optErr) e') -| _ -> ret (Obj.magic coq_Monad_optErr) set_empty - -(** val get_use_exprs : expr list -> coq_LVDomain optErr **) - -let rec get_use_exprs = function -| Coq_nil -> ret (Obj.magic coq_Monad_optErr) set_empty -| Coq_cons (e, rest) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr e) (fun e' -> - bind (Obj.magic coq_Monad_optErr) (get_use_exprs rest) (fun rest' -> - ret (Obj.magic coq_Monad_optErr) - (set_union (app (PTree.xkeys rest' Coq_xH) (PTree.xkeys e' Coq_xH)) - rest' e' PTree.empty))) - -(** val get_use : StmtCGraph.statement -> coq_LVDomain optErr **) - -let get_use = function -| StmtCGraph.Smassign (lv, rv, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr lv) (fun lv' -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr rv) (fun rv' -> - ret (Obj.magic coq_Monad_optErr) - (set_union (app (PTree.xkeys lv' Coq_xH) (PTree.xkeys rv' Coq_xH)) - lv' rv' PTree.empty))) -| StmtCGraph.Ssassign (lv, rv, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr lv) (fun lv' -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr rv) (fun rv' -> - ret (Obj.magic coq_Monad_optErr) - (set_union (app (PTree.xkeys lv' Coq_xH) (PTree.xkeys rv' Coq_xH)) - lv' rv' PTree.empty))) -| StmtCGraph.Sset (_, rv, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr rv) (fun rv' -> - ret (Obj.magic coq_Monad_optErr) rv') -| StmtCGraph.Scall (_, _, args, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_exprs args) (fun args' -> - ret (Obj.magic coq_Monad_optErr) args') -| Scond (cond, _, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr cond) (fun cond' -> - ret (Obj.magic coq_Monad_optErr) cond') -| StmtCGraph.Sreturn (val0, _) -> - (match val0 with - | Some e -> ret (Obj.magic coq_Monad_optErr) (set_add e set_empty) - | None -> ret (Obj.magic coq_Monad_optErr) set_empty) -| StmtCGraph.Shash (e1, e2, eo, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr e1) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr e2) (fun e2' -> - match eo with - | Some eoo -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr eoo) (fun eoo' -> - ret (Obj.magic coq_Monad_optErr) - (set_union - (app - (PTree.xkeys - (set_union - (app (PTree.xkeys e1' Coq_xH) (PTree.xkeys e2' Coq_xH)) - e1' e2' PTree.empty) Coq_xH) (PTree.xkeys eoo' Coq_xH)) - (set_union - (app (PTree.xkeys e1' Coq_xH) (PTree.xkeys e2' Coq_xH)) e1' - e2' PTree.empty) eoo' PTree.empty)) - | None -> - ret (Obj.magic coq_Monad_optErr) - (set_union (app (PTree.xkeys e1' Coq_xH) (PTree.xkeys e2' Coq_xH)) - e1' e2' PTree.empty))) -| StmtCGraph.Stransfer (a, v, _, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr a) (fun a' -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr v) (fun v' -> - ret (Obj.magic coq_Monad_optErr) - (set_union (app (PTree.xkeys a' Coq_xH) (PTree.xkeys v' Coq_xH)) a' - v' PTree.empty))) -| StmtCGraph.Scallmethod (a, _, _, v, args, _, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr a) (fun a' -> - bind (Obj.magic coq_Monad_optErr) (get_use_expr v) (fun v' -> - bind (Obj.magic coq_Monad_optErr) (get_use_exprs args) (fun args' -> - ret (Obj.magic coq_Monad_optErr) - (set_union - (app - (PTree.xkeys - (set_union - (app (PTree.xkeys a' Coq_xH) (PTree.xkeys v' Coq_xH)) a' v' - PTree.empty) Coq_xH) (PTree.xkeys args' Coq_xH)) - (set_union (app (PTree.xkeys a' Coq_xH) (PTree.xkeys v' Coq_xH)) - a' v' PTree.empty) args' PTree.empty)))) -| StmtCGraph.Slog (topics, args, _) -> - bind (Obj.magic coq_Monad_optErr) (get_use_exprs topics) (fun topics' -> - bind (Obj.magic coq_Monad_optErr) (get_use_exprs args) (fun args' -> - ret (Obj.magic coq_Monad_optErr) - (set_union - (app (PTree.xkeys topics' Coq_xH) (PTree.xkeys args' Coq_xH)) - topics' args' PTree.empty))) -| _ -> ret (Obj.magic coq_Monad_optErr) set_empty - -(** val get_successor : StmtCGraph.statement -> set **) - -let get_successor = function -| StmtCGraph.Sskip n -> set_add n set_empty -| StmtCGraph.Smassign (_, _, n) -> set_add n set_empty -| StmtCGraph.Ssassign (_, _, n) -> set_add n set_empty -| StmtCGraph.Sset (_, _, n) -> set_add n set_empty -| StmtCGraph.Scall (_, _, _, n) -> set_add n set_empty -| Scond (_, ntrue, nfalse) -> - set_union - (app (PTree.xkeys (set_add ntrue set_empty) Coq_xH) - (PTree.xkeys (set_add nfalse set_empty) Coq_xH)) - (set_add ntrue set_empty) (set_add nfalse set_empty) PTree.empty -| StmtCGraph.Sreturn (_, n) -> set_add n set_empty -| StmtCGraph.Shash (_, _, _, n) -> set_add n set_empty -| StmtCGraph.Stransfer (_, _, nfail, n) -> - set_union - (app (PTree.xkeys (set_add nfail set_empty) Coq_xH) - (PTree.xkeys (set_add n set_empty) Coq_xH)) (set_add nfail set_empty) - (set_add n set_empty) PTree.empty -| StmtCGraph.Scallmethod (_, _, _, _, _, nfail, n) -> - set_union - (app (PTree.xkeys (set_add nfail set_empty) Coq_xH) - (PTree.xkeys (set_add n set_empty) Coq_xH)) (set_add nfail set_empty) - (set_add n set_empty) PTree.empty -| StmtCGraph.Slog (_, _, n) -> set_add n set_empty -| _ -> set_empty - -(** val use : code -> node -> coq_LVDomain optErr **) - -let use cfg block = - match PTree.get block cfg with - | Some s -> - bind (Obj.magic coq_Monad_optErr) (get_use s) (fun s' -> - ret (Obj.magic coq_Monad_optErr) s') - | None -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - -(** val def : code -> node -> coq_LVDomain optErr **) - -let def cfg block = - match PTree.get block cfg with - | Some s -> - (match s with - | StmtCGraph.Sset (id, _, _) -> - ret (Obj.magic coq_Monad_optErr) (set_add id set_empty) - | StmtCGraph.Scall (retval, _, _, _) -> - (match retval with - | Some rv -> ret (Obj.magic coq_Monad_optErr) (set_add rv set_empty) - | None -> ret (Obj.magic coq_Monad_optErr) set_empty) - | StmtCGraph.Scallmethod (_, rvs, _, _, _, _, _) -> - ret (Obj.magic coq_Monad_optErr) - (fold_left (fun acc e -> set_add e acc) rvs set_empty) - | _ -> ret (Obj.magic coq_Monad_optErr) set_empty) - | None -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - -(** val f_b : code -> node -> coq_LVDomain -> coq_LVDomain optErr **) - -let f_b cfg block x = - bind (Obj.magic coq_Monad_optErr) (use cfg block) (fun u -> - bind (Obj.magic coq_Monad_optErr) (def cfg block) (fun d -> - ret (Obj.magic coq_Monad_optErr) - (set_union - (app (PTree.xkeys u Coq_xH) - (PTree.xkeys (set_minus (PTree.xkeys x Coq_xH) x d PTree.empty) - Coq_xH)) u (set_minus (PTree.xkeys x Coq_xH) x d PTree.empty) - PTree.empty))) - -type coq_CD = coq_LVDomain PTree.t - -(** val initialize : nat -> coq_CD -> coq_CD **) - -let rec initialize i s = - match i with - | O -> s - | S x -> PTree.set (Pos.of_nat x) set_empty s - -(** val empty_domain : set **) - -let empty_domain = - set_empty - -(** val get_lv : - code -> node -> node -> nat -> coq_CD -> coq_CD -> (coq_CD, coq_CD) prod - optErr **) - -let rec get_lv cfg nret nrev i iN oUT = - let id = Pos.of_nat i in - (match is_exit nret nrev id with - | Coq_true -> - (match i with - | O -> ret (Obj.magic coq_Monad_optErr) (Coq_pair (iN, oUT)) - | S x -> get_lv cfg nret nrev x iN oUT) - | Coq_false -> - (match PTree.get id cfg with - | Some stmt -> - let oUT0 = - PTree.set id - (set_fold_left (fun acc e -> - set_union - (app (PTree.xkeys acc Coq_xH) - (PTree.xkeys (PTree.get_default set_empty e iN) Coq_xH)) - acc (PTree.get_default set_empty e iN) PTree.empty) - (get_successor stmt) empty_domain) oUT - in - bind (Obj.magic coq_Monad_optErr) - (Obj.magic f_b cfg id (PTree.get_default set_empty id oUT0)) - (fun v -> - let iN0 = PTree.set id v iN in - (match i with - | O -> ret (Obj.magic coq_Monad_optErr) (Coq_pair (iN0, oUT0)) - | S x -> get_lv cfg nret nrev x iN0 oUT0)) - | None -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))) - -(** val get_lv_mfp : - code -> nat -> node -> node -> nat -> coq_CD -> coq_CD -> coq_CD optErr **) - -let rec get_lv_mfp cfg maxnode nret nrev iteration iN oUT = - match iteration with - | O -> ret (Obj.magic coq_Monad_optErr) oUT - | S x -> - (match get_lv cfg nret nrev maxnode iN oUT with - | Success p -> - let Coq_pair (nIN, nOUT) = p in - (match PTree.beq set_eq iN nIN with - | Coq_true -> ret (Obj.magic coq_Monad_optErr) oUT - | Coq_false -> get_lv_mfp cfg maxnode nret nrev x nIN nOUT) - | Error msg -> Error msg) - -(** val compute_livevar : - code -> nat -> node -> node -> nat -> coq_CD optErr **) - -let compute_livevar cfg maxnode nret nrev precision = - let iN = initialize maxnode PTree.empty in - let oUT = initialize maxnode PTree.empty in - get_lv_mfp cfg maxnode nret nrev precision iN oUT - -type tvs = set - -type clashg = tvs PTree.t - -(** val create_clash_graph : coq_CD -> clashg optErr **) - -let rec create_clash_graph lv = - ret (Obj.magic coq_Monad_optErr) - (PTree.fold1 (fun acc e -> - set_fold_left (fun acc' e' -> - PTree.set e' - (set_union - (app (PTree.xkeys (PTree.get_default set_empty e' acc') Coq_xH) - (PTree.xkeys e Coq_xH)) (PTree.get_default set_empty e' acc') e - PTree.empty) acc') e acc) lv PTree.empty) - -(** val remove_node_graph : ident -> clashg -> clashg **) - -let rec remove_node_graph t0 g = - let g' = PTree.remove t0 g in - PTree.fold (fun acc i e -> - PTree.set i - (set_minus (PTree.xkeys e Coq_xH) e (set_add t0 set_empty) PTree.empty) - acc) g' g' - -(** val least_degree : clashg -> ident **) - -let rec least_degree g = - let Coq_pair (id, _) = - PTree.fold (fun acc i e -> - let Coq_pair (_, mn) = acc in - let c = set_card e in - (match Nat.leb c mn with - | Coq_true -> Coq_pair (i, c) - | Coq_false -> acc)) g (Coq_pair ((Pos.of_nat O), (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S O)))))))))))))))))) - in - id - -(** val pop_graph : nat -> clashg -> ident list -> ident list optErr **) - -let rec pop_graph i g stack = - match i with - | O -> - (match Nat.eqb (length (PTree.elements g)) O with - | Coq_true -> ret (Obj.magic coq_Monad_optErr) stack - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | S n -> - let ldn = least_degree g in - pop_graph n (remove_node_graph ldn g) (Coq_cons (ldn, stack)) - -type colorMap = nat PTree.t - -(** val get_assigned_colors : colorMap -> tvs -> set **) - -let get_assigned_colors c s = - PTree.fold (fun acc i e -> - match e with - | Coq_true -> - (match PTree.get i c with - | Some cl -> set_add (Pos.of_nat cl) acc - | None -> acc) - | Coq_false -> acc) s set_empty - -(** val colorMapFull : set **) - -let colorMapFull = - set_add (Pos.of_nat (S O)) - (set_add (Pos.of_nat (S (S O))) - (set_add (Pos.of_nat (S (S (S O)))) - (set_add (Pos.of_nat (S (S (S (S O))))) - (set_add (Pos.of_nat (S (S (S (S (S O)))))) - (set_add (Pos.of_nat (S (S (S (S (S (S O))))))) - (set_add (Pos.of_nat (S (S (S (S (S (S (S O)))))))) - (set_add (Pos.of_nat (S (S (S (S (S (S (S (S O))))))))) - (set_add (Pos.of_nat (S (S (S (S (S (S (S (S (S O)))))))))) - (set_add - (Pos.of_nat (S (S (S (S (S (S (S (S (S (S O))))))))))) - (set_add - (Pos.of_nat (S (S (S (S (S (S (S (S (S (S (S - O)))))))))))) - (set_add - (Pos.of_nat (S (S (S (S (S (S (S (S (S (S (S (S - O))))))))))))) - (set_add - (Pos.of_nat (S (S (S (S (S (S (S (S (S (S (S (S - (S O)))))))))))))) - (set_add - (Pos.of_nat (S (S (S (S (S (S (S (S (S (S (S (S - (S (S O))))))))))))))) - (set_add - (Pos.of_nat (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S O)))))))))))))))) set_empty)))))))))))))) - -(** val max_reg : nat **) - -let max_reg = - S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S O))))))))))))))) - -(** val default_reg : nat **) - -let default_reg = - S O - -(** val assign_color : ident list -> clashg -> colorMap -> colorMap optErr **) - -let rec assign_color stack g c = - match stack with - | Coq_nil -> ret (Obj.magic coq_Monad_optErr) c - | Coq_cons (x, xs) -> - (match PTree.get x g with - | Some lk -> - let assignable = - set_minus (PTree.xkeys colorMapFull Coq_xH) colorMapFull - (get_assigned_colors c lk) PTree.empty - in - (match Nat.ltb O (set_card assignable) with - | Coq_true -> - (match PTree.fold (fun acc i e -> - match e with - | Coq_true -> Some i - | Coq_false -> acc) assignable None with - | Some reg -> assign_color xs g (PTree.set x (Pos.to_nat reg) c) - | None -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))) - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | None -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - -(** val recolor_expr : expr -> colorMap -> expr optErr **) - -let rec recolor_expr e c = - match e with - | Etempvar (i, t0) -> - ret (Obj.magic coq_Monad_optErr) (Etempvar - ((Pos.of_nat (PTree.get_default default_reg i c)), t0)) - | Esload (e0, t0) -> - bind (Obj.magic coq_Monad_optErr) (recolor_expr e0 c) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (Esload (e', t0))) - | Emload (e0, t0) -> - bind (Obj.magic coq_Monad_optErr) (recolor_expr e0 c) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (Emload (e', t0))) - | Eaddr (e0, t0) -> - bind (Obj.magic coq_Monad_optErr) (recolor_expr e0 c) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (Eaddr (e', t0))) - | Eunop (o, e0, t0) -> - bind (Obj.magic coq_Monad_optErr) (recolor_expr e0 c) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (Eunop (o, e', t0))) - | Ebinop (o, e1, e2, t0) -> - bind (Obj.magic coq_Monad_optErr) (recolor_expr e1 c) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) (recolor_expr e2 c) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) (Ebinop (o, e1', e2', t0)))) - | Ecall1 (b, e0, t0) -> - bind (Obj.magic coq_Monad_optErr) (recolor_expr e0 c) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (Ecall1 (b, e', t0))) - | _ -> ret (Obj.magic coq_Monad_optErr) e - -(** val recolor_exprs : expr list -> colorMap -> expr list optErr **) - -let rec recolor_exprs es c = - match es with - | Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil - | Coq_cons (e, rest) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr e c) (fun e' -> - bind (Obj.magic coq_Monad_optErr) (recolor_exprs rest c) (fun rest' -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (e', rest')))) - -(** val recolor_regs : - ident list -> colorMap -> ident list -> ident list optErr **) - -let rec recolor_regs l c r = - match l with - | Coq_nil -> ret (Obj.magic coq_Monad_optErr) r - | Coq_cons (x, xs) -> - recolor_regs xs c (Coq_cons - ((Pos.of_nat (PTree.get_default default_reg x c)), r)) - -(** val recolor_reg : ident -> colorMap -> ident **) - -let recolor_reg i c = - Pos.of_nat (PTree.get_default default_reg i c) - -(** val recolor : - code -> set -> nat -> node -> colorMap -> (code -> code) -> (code -> - code) optErr **) - -let rec recolor cfg t0 rd tn c cont = - match rd with - | O -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | S x -> - (match PTree.get tn cfg with - | Some s -> - (match set_in tn t0 with - | Coq_true -> ret (Obj.magic coq_Monad_optErr) cont - | Coq_false -> - let t' = - set_union - (app (PTree.xkeys t0 Coq_xH) - (PTree.xkeys (set_add tn set_empty) Coq_xH)) t0 - (set_add tn set_empty) PTree.empty - in - let recolor' = recolor cfg t' x in - (match s with - | StmtCGraph.Sskip n -> recolor' n c cont - | StmtCGraph.Smassign (lv, rv, n) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr lv c) - (fun lv' -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_expr rv c) (fun rv' -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Smassign (lv', rv', n)) (cont g)))) - | StmtCGraph.Ssassign (lv, rv, n) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr lv c) - (fun lv' -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_expr rv c) (fun rv' -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Ssassign (lv', rv', n)) (cont g)))) - | StmtCGraph.Sset (id, rv, n) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr rv c) - (fun rv' -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Sset ((recolor_reg id c), rv', n)) - (cont g))) - | StmtCGraph.Scall (retval, lab, args, n) -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_exprs args c) (fun args' -> - match retval with - | Some rv -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Scall ((Some (recolor_reg rv c)), - lab, args', n)) (cont g)) - | None -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Scall (retval, lab, args', n)) - (cont g))) - | Scond (cond, ntrue, nfalse) -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_expr cond c) (fun cond' -> - bind (Obj.magic coq_Monad_optErr) - (recolor' ntrue c (fun g -> g)) (fun cont' -> - recolor' nfalse c (fun g -> - PTree.set tn (Scond (cond', ntrue, nfalse)) - (cont' (cont g))))) - | StmtCGraph.Sreturn (val0, n) -> - (match val0 with - | Some v -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Sreturn ((Some (recolor_reg v c)), - n)) (cont g)) - | None -> recolor' n c cont) - | StmtCGraph.Shash (e1, e2, eo, n) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr e1 c) - (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_expr e2 c) (fun e2' -> - match eo with - | Some eoo -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_expr eoo c) (fun eoo' -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Shash (e1', e2', (Some eoo'), - n)) (cont g))) - | None -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Shash (e1', e2', None, n)) - (cont g)))) - | StmtCGraph.Stransfer (a, v, nfail, n) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr a c) - (fun a' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr v c) - (fun v' -> - bind (Obj.magic coq_Monad_optErr) - (recolor' nfail c (fun g -> g)) (fun cont' -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Stransfer (a', v', nfail, n)) - (cont' (cont g)))))) - | StmtCGraph.Scallmethod (a, rvs, sig0, v, args, nfail, n) -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_regs rvs c Coq_nil) (fun rvs' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr a c) - (fun a' -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_expr v c) (fun v' -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_exprs args c) (fun args' -> - bind (Obj.magic coq_Monad_optErr) - (recolor' nfail c (fun g -> g)) (fun cont' -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Scallmethod (a', rvs', - sig0, v', args', nfail, n)) (cont' (cont g)))))))) - | StmtCGraph.Slog (topics, args, n) -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_exprs topics c) (fun topics' -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor_exprs args c) (fun args' -> - recolor' n c (fun g -> - PTree.set tn (StmtCGraph.Slog (topics', args', n)) (cont g)))) - | _ -> ret (Obj.magic coq_Monad_optErr) cont)) - | None -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - -(** val color_graph : - code -> nat -> node -> clashg -> (code, colorMap) prod optErr **) - -let color_graph cfg maxnode nentry cg = - match PTree.fold (fun acc _ e -> - match acc with - | Coq_true -> Nat.ltb (set_card e) max_reg - | Coq_false -> Coq_false) cg Coq_true with - | Coq_true -> - let num_tv = length (PTree.elements cg) in - bind (Obj.magic coq_Monad_optErr) - (Obj.magic pop_graph (length (PTree.elements cg)) cg Coq_nil) - (fun stack -> - let unique_nodes = - Nat.eqb - (set_card (fold_left (fun acc e -> set_add e acc) stack set_empty)) - num_tv - in - (match unique_nodes with - | Coq_true -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic assign_color stack cg PTree.empty) (fun colors -> - match Nat.eqb (length (PTree.elements colors)) num_tv with - | Coq_true -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic recolor cfg set_empty maxnode nentry colors - (fun g -> g)) (fun f -> - ret (Obj.magic coq_Monad_optErr) (Coq_pair ((f cfg), colors))) - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Coq_false -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - -(** val variable_coalescing : - code -> nat -> node -> node -> node -> nat -> (code, colorMap) prod optErr **) - -let variable_coalescing cfg maxnode nentry nret nrev precision = - bind (Obj.magic coq_Monad_optErr) - (Obj.magic compute_livevar cfg maxnode nret nrev precision) (fun lvs -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic create_clash_graph lvs) - (fun cg -> color_graph cfg maxnode nentry cg)) - -(** val get_clash_graph : - code -> nat -> node -> node -> nat -> (positive, positive list) prod list - optErr **) - -let get_clash_graph cfg maxnode nret nrev precision = - bind (Obj.magic coq_Monad_optErr) - (Obj.magic compute_livevar cfg maxnode nret nrev precision) (fun lvs -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic create_clash_graph lvs) - (fun cg -> - ret (Obj.magic coq_Monad_optErr) - (map (fun x -> - let Coq_pair (id, tvss) = x in - Coq_pair (id, - (PTree.fold (fun acc i ind -> - match ind with - | Coq_true -> Coq_cons (i, acc) - | Coq_false -> acc) tvss Coq_nil))) (PTree.elements cg)))) - -(** val max_iteration : nat **) - -let max_iteration = - S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S - O)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - (** val cgraph_function : coq_function -> StmtCGraph.coq_function optErr **) let cgraph_function f = @@ -2046,82 +175,10 @@ let cgraph_function f = Coq_true, Coq_false)), EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))) | OK (nentry, s) -> - (match variable_coalescing s.st_code - (sub (Pos.to_nat s.st_nextnode) (S O)) nentry - (Pos.of_nat (S O)) (Pos.of_nat (S (S O))) max_iteration with - | Success p -> - let Coq_pair (ncfg, cmap) = p in - ret (Obj.magic coq_Monad_optErr) { StmtCGraph.fn_return = - f.fn_return; StmtCGraph.fn_params = - (map (fun e -> - let Coq_pair (id, ty) = e in - Coq_pair ((Pos.of_nat (PTree.get_default (S O) id cmap)), ty)) - f.fn_params); StmtCGraph.fn_temps = - (map (fun e -> - let Coq_pair (id, ty) = e in - Coq_pair ((Pos.of_nat (PTree.get_default (S O) id cmap)), ty)) - f.fn_temps); StmtCGraph.fn_locals = f.fn_locals; fn_code = ncfg; - fn_entrypoint = nentry } - | Error msg -> Error msg)) - -(** val clash_function : - coq_function -> (positive, positive list) prod list optErr **) - -let clash_function f = - let cgraph_fun = - bind (Obj.magic coq_Monad_mon) (add_instr Sdone) (fun nret -> - bind (Obj.magic coq_Monad_mon) (add_instr StmtCGraph.Srevert) - (fun nrev -> cgraph_statement f.fn_body nret nret nrev None)) - in - (match cgraph_fun init_state with - | Fail -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))) - | OK (_, s) -> - get_clash_graph s.st_code (sub (Pos.to_nat s.st_nextnode) (S O)) - (Pos.of_nat (S O)) (Pos.of_nat (S (S O))) max_iteration) + ret (Obj.magic coq_Monad_optErr) { StmtCGraph.fn_return = f.fn_return; + StmtCGraph.fn_params = f.fn_params; StmtCGraph.fn_temps = f.fn_temps; + StmtCGraph.fn_locals = f.fn_locals; fn_code = s.st_code; + fn_entrypoint = nentry; fn_nextnode = s.st_nextnode }) (** val empty_constructor : coq_function **) @@ -2154,38 +211,6 @@ let cgraph_functions defs = let cgraph_methoddefs defs = transl_map cgraph_function defs -(** val graphviz_helper : - StmtCGraph.coq_function -> ((positive, StmtCGraph.statement) prod list, - positive) prod **) - -let graphviz_helper f = - Coq_pair ((PTree.elements f.fn_code), f.fn_entrypoint) - -(** val clash_viz : - genv -> (positive, positive list) prod list list optErr **) - -let clash_viz ge = - fold_left (fun acc x -> - match acc with - | Success acc' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic clash_function x) - (fun x' -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (x', acc'))) - | Error msg -> Error msg) (Genv.all_functions ge) (Success Coq_nil) - -(** val cgraph_viz : - genv -> ((positive, StmtCGraph.statement) prod list, positive) prod list - optErr **) - -let cgraph_viz ge = - fold_left (fun acc x -> - match acc with - | Success acc' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cgraph_function x) - (fun x' -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((graphviz_helper x'), - acc'))) - | Error msg -> Error msg) (Genv.all_functions ge) (Success Coq_nil) - (** val cgraph_genv : genv -> StmtCGraph.genv optErr **) let cgraph_genv ge = diff --git a/src/backend/extraction/Gen1.mli b/src/backend/extraction/Gen1.mli index 3d4b720..de8b7a4 100644 --- a/src/backend/extraction/Gen1.mli +++ b/src/backend/extraction/Gen1.mli @@ -1,18 +1,13 @@ -open AST open Ascii open BinNums open BinPos open Coqlib open Ctypes open Datatypes -open ExpCintptr open Globalenvs -open List0 open Maps0 open Monad -open Nat0 open OptErrMonad -open PeanoNat open Specif open StmtCGraph open StmtCintptr @@ -50,110 +45,8 @@ val update_instr : node -> StmtCGraph.statement -> coq_unit mon val cgraph_statement : statement -> node -> node -> node -> node option -> node mon -type set = bool PTree.t - -type coq_LVDomain = set - -val set_empty : set - -val set_union : positive list -> set -> set -> set -> set - -val set_minus : positive list -> set -> set -> set -> set - -val set_add : ident -> set -> set - -val set_in : ident -> set -> bool - -val set_card : set -> nat - -val set_eq : set -> set -> bool - -val set_fold_left : ('a1 -> ident -> 'a1) -> set -> 'a1 -> 'a1 - -val is_exit : node -> node -> node -> bool - -val get_use_expr : expr -> coq_LVDomain optErr - -val get_use_exprs : expr list -> coq_LVDomain optErr - -val get_use : StmtCGraph.statement -> coq_LVDomain optErr - -val get_successor : StmtCGraph.statement -> set - -val use : code -> node -> coq_LVDomain optErr - -val def : code -> node -> coq_LVDomain optErr - -val f_b : code -> node -> coq_LVDomain -> coq_LVDomain optErr - -type coq_CD = coq_LVDomain PTree.t - -val initialize : nat -> coq_CD -> coq_CD - -val empty_domain : set - -val get_lv : - code -> node -> node -> nat -> coq_CD -> coq_CD -> (coq_CD, coq_CD) prod - optErr - -val get_lv_mfp : - code -> nat -> node -> node -> nat -> coq_CD -> coq_CD -> coq_CD optErr - -val compute_livevar : code -> nat -> node -> node -> nat -> coq_CD optErr - -type tvs = set - -type clashg = tvs PTree.t - -val create_clash_graph : coq_CD -> clashg optErr - -val remove_node_graph : ident -> clashg -> clashg - -val least_degree : clashg -> ident - -val pop_graph : nat -> clashg -> ident list -> ident list optErr - -type colorMap = nat PTree.t - -val get_assigned_colors : colorMap -> tvs -> set - -val colorMapFull : set - -val max_reg : nat - -val default_reg : nat - -val assign_color : ident list -> clashg -> colorMap -> colorMap optErr - -val recolor_expr : expr -> colorMap -> expr optErr - -val recolor_exprs : expr list -> colorMap -> expr list optErr - -val recolor_regs : ident list -> colorMap -> ident list -> ident list optErr - -val recolor_reg : ident -> colorMap -> ident - -val recolor : - code -> set -> nat -> node -> colorMap -> (code -> code) -> (code -> code) - optErr - -val color_graph : - code -> nat -> node -> clashg -> (code, colorMap) prod optErr - -val variable_coalescing : - code -> nat -> node -> node -> node -> nat -> (code, colorMap) prod optErr - -val get_clash_graph : - code -> nat -> node -> node -> nat -> (positive, positive list) prod list - optErr - -val max_iteration : nat - val cgraph_function : coq_function -> StmtCGraph.coq_function optErr -val clash_function : - coq_function -> (positive, positive list) prod list optErr - val empty_constructor : coq_function val cgraph_constructor : @@ -166,14 +59,4 @@ val cgraph_methoddefs : coq_function option IntMap.t -> StmtCGraph.coq_function option IntMap.t optErr -val graphviz_helper : - StmtCGraph.coq_function -> ((positive, StmtCGraph.statement) prod list, - positive) prod - -val clash_viz : genv -> (positive, positive list) prod list list optErr - -val cgraph_viz : - genv -> ((positive, StmtCGraph.statement) prod list, positive) prod list - optErr - val cgraph_genv : genv -> StmtCGraph.genv optErr diff --git a/src/backend/extraction/Gen10.ml b/src/backend/extraction/Gen10.ml new file mode 100644 index 0000000..68a38a6 --- /dev/null +++ b/src/backend/extraction/Gen10.ml @@ -0,0 +1,1913 @@ +open AST +open Ascii +open BinNums +open BinPos +open Cop +open Ctypes +open Datatypes +open ExpCintptr +open GlobalenvCompile +open Globalenvs +open Int0 +open Integers +open Language0 +open List0 +open Maps0 +open Monad +open Nat0 +open OptErrMonad +open Options +open PeanoNat +open StackEnv +open StmtCintptr +open String0 +open Structure +open Trees +open Values + +let __ = let rec f _ = Obj.repr f in Obj.repr f + +(** val wasm_expr : + nat PTree.t -> ExpCintptr.expr -> bool -> instr list optErr **) + +let rec wasm_expr temps e rvalue = + match e with + | Econst_int (_, _) -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Econst_int256 (i, _) -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Const_256 i), Coq_nil)) + | Etempvar (i, _) -> + (match rvalue with + | Coq_true -> + bind (Obj.magic coq_Monad_optErr) + (fromOption (PTree.get i (Obj.magic temps)) (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + (fun ind -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Local_get ind), + Coq_nil))) + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Esload (e0, _) -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e0 Coq_false) + (fun e' -> + ret (Obj.magic coq_Monad_optErr) + (app e' + (app + (set_eeiOffset Coq_nil Coq_pathOffset (S (S (S (S (S (S (S (S + O))))))))) + (app (Coq_cons ((Const_nat + (get_idx_eei_mem_offset Coq_pathOffset)), (Coq_cons ((Const_nat + (get_idx_eei_mem_offset Coq_resultOffset)), (Coq_cons ((Call + (get_idx_eei Coq_eei_storageLoad)), Coq_nil)))))) + (load_resultOffset (S O)))))) + | Emload (e0, t0) -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e0 Coq_false) + (fun e' -> + bind (Obj.magic coq_Monad_optErr) + (match t0 with + | Tpointer (p, _) -> + (match p with + | Coq_stor -> + Success (Coq_cons ((Const_nat (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S O))))))))))))))))))))))))))))))))), (Coq_cons ((Binop + (Coq_i32 IOp32.Add)), (Coq_cons ((Global_set + scratch_global_idx2), + (app Coq_nil + (load_len_hash (S (S (S (S (S (S (S (S O)))))))))))))))) + | _ -> Success (Coq_cons ((Memop (Coq_i32 IOp32.Load)), Coq_nil))) + | _ -> Success (Coq_cons ((Memop (Coq_i32 IOp32.Load)), Coq_nil))) + (fun l -> ret (Obj.magic coq_Monad_optErr) (app e' l))) + | Eaddr (e0, _) -> + (match rvalue with + | Coq_true -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e0 Coq_false) + (fun e' -> ret (Obj.magic coq_Monad_optErr) e') + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Eunop (o, e0, _) -> + (match rvalue with + | Coq_true -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e0 Coq_true) + (fun e' -> + bind (Obj.magic coq_Monad_optErr) + (match o with + | Onotbool -> Success (Coq_cons ((Testop (Coq_i32 __)), Coq_nil)) + | Onotint -> + Success (Coq_cons ((Call (get_idx_aux Coq_aux_notint)), + Coq_nil)) + | Oneg -> + Success (Coq_cons ((Const (Coq_i32 I32.zero)), (Coq_cons + ((Binop (Coq_i32 IOp32.Sub)), Coq_nil)))) + | Osha_1 -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + (fun o' -> ret (Obj.magic coq_Monad_optErr) (app e' o'))) + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Ebinop (o, e1, e2, _) -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e1 rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e2 Coq_true) + (fun e2' -> + bind (Obj.magic coq_Monad_optErr) + (match o with + | Oadd -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Add)), Coq_nil)) + | Osub -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Sub)), Coq_nil)) + | Omul -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Mul)), Coq_nil)) + | Odiv -> + Success (Coq_cons ((Binop (Coq_i32 (IOp32.Div Unsigned))), + Coq_nil)) + | Omod -> + Success (Coq_cons ((Binop (Coq_i32 (IOp32.Rem Unsigned))), + Coq_nil)) + | Oexp -> + Success (Coq_cons ((Call (get_idx_aux Coq_aux_pow)), Coq_nil)) + | Oand -> Success (Coq_cons ((Binop (Coq_i32 IOp32.And)), Coq_nil)) + | Oor -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Or)), Coq_nil)) + | Oxor -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Xor)), Coq_nil)) + | Oshl -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Shl)), Coq_nil)) + | Oshr -> + Success (Coq_cons ((Binop (Coq_i32 (IOp32.Shr Unsigned))), + Coq_nil)) + | Oeq -> Success (Coq_cons ((Relop (Coq_i32 IOp32.Eq)), Coq_nil)) + | One -> Success (Coq_cons ((Relop (Coq_i32 IOp32.Ne)), Coq_nil)) + | Olt -> + Success (Coq_cons ((Relop (Coq_i32 (IOp32.Lt Unsigned))), + Coq_nil)) + | Ogt -> + Success (Coq_cons ((Relop (Coq_i32 (IOp32.Gt Unsigned))), + Coq_nil)) + | Ole -> + Success (Coq_cons ((Relop (Coq_i32 (IOp32.Le Unsigned))), + Coq_nil)) + | Oge -> + Success (Coq_cons ((Relop (Coq_i32 (IOp32.Ge Unsigned))), + Coq_nil)) + | Osha_2 -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + (fun o' -> ret (Obj.magic coq_Monad_optErr) (app e1' (app e2' o'))))) + | Ecall0 (b, _) -> + (match rvalue with + | Coq_true -> + bind (Obj.magic coq_Monad_optErr) (wasm_builtin0 b) (fun b' -> + ret (Obj.magic coq_Monad_optErr) b') + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Ecall1 (b, e0, _) -> + (match rvalue with + | Coq_true -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e0 Coq_true) + (fun e' -> + bind (Obj.magic coq_Monad_optErr) (Success (wasm_builtin1 b)) + (fun b' -> ret (Obj.magic coq_Monad_optErr) (app e' b'))) + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + +(** val wasm_exprs : + nat PTree.t -> ExpCintptr.expr list -> instr list optErr **) + +let rec wasm_exprs temps = function +| Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil +| Coq_cons (e, rest) -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e Coq_true) (fun e' -> + bind (Obj.magic coq_Monad_optErr) (wasm_exprs temps rest) (fun rest' -> + ret (Obj.magic coq_Monad_optErr) (app rest' e'))) + +(** val optident : nat PTree.t -> ident option -> instrs optErr **) + +let optident temps = function +| Some b -> + (match PTree.get b temps with + | Some v -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Local_set v), Coq_nil)) + | None -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))) +| None -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (Drop, Coq_nil)) + +(** val wasm_statement : + (positive, nat) prod list -> nat PTree.t -> coq_type -> statement -> nat + -> instrs optErr **) + +let rec wasm_statement funident_map temps return_type s cstack_depth = + match s with + | Sskip -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (Nop, Coq_nil)) + | Ssassign (lv, rv) -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps lv Coq_false) + (fun lv' -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps rv Coq_true) + (fun rv' -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app + (set_eeiOffset Coq_nil Coq_pathOffset (S (S (S (S (S (S (S (S + O))))))))) + (app (set_eeiOffset rv' Coq_valueOffset (S O)) (Coq_cons + ((Const_nat (get_idx_eei_mem_offset Coq_pathOffset)), + (Coq_cons ((Const_nat + (get_idx_eei_mem_offset Coq_valueOffset)), (Coq_cons ((Call + (get_idx_eei Coq_eei_storageStore)), Coq_nil))))))))))) + | Smassign (lv, rv) -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps lv Coq_false) + (fun lv' -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps rv Coq_true) + (fun rv' -> + match typeof lv with + | Tpointer (_, _) -> + (match rv with + | Econst_int (_, _) -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) + | Econst_int256 (_, _) -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) + | Etempvar (_, _) -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) + | Esload (_, _) -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) + | Emload (_, t0) -> + (match t0 with + | Tpointer (mem, _) -> + (match mem with + | Coq_mem -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), + Coq_nil)))) + | Coq_stor -> + ret (Obj.magic coq_Monad_optErr) + (app rv' + (store_len lv' (S (S (S (S (S (S (S (S O)))))))))) + | Coq_call -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), + Coq_nil))))) + | _ -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Eaddr (_, _) -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) + | Eunop (_, _, _) -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) + | Ebinop (_, _, _, _) -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) + | Ecall0 (_, _) -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) + | Ecall1 (_, _, _) -> + ret (Obj.magic coq_Monad_optErr) + (app lv' + (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil))))) + | _ -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Sset (id, rv) -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps rv Coq_true) + (fun rv' -> + bind (Obj.magic coq_Monad_optErr) + (match rv with + | Emload (_, t0) -> + (match t0 with + | Tpointer (p, _) -> + (match p with + | Coq_stor -> + Success (Coq_cons (Drop, (Coq_cons (Drop, (Coq_cons (Drop, + (Coq_cons (Drop, (Coq_cons (Drop, (Coq_cons (Drop, + (Coq_cons (Drop, Coq_nil)))))))))))))) + | _ -> Success Coq_nil) + | _ -> Success Coq_nil) + | _ -> Success Coq_nil) (fun d -> + let set_lookup = + match PTree.get id temps with + | Some v -> Coq_cons ((Local_set v), Coq_nil) + | None -> Coq_nil + in + ret (Obj.magic coq_Monad_optErr) (app rv' (app d set_lookup)))) + | Scall (retval, lab, args) -> + bind (Obj.magic coq_Monad_optErr) (optident temps retval) (fun t0 -> + bind (Obj.magic coq_Monad_optErr) (wasm_exprs temps args) (fun args' -> + let ido = + fold_left (fun xs x -> + match xs with + | Some _ -> xs + | None -> + (match Pos.eqb (fst x) lab with + | Coq_true -> Some (snd x) + | Coq_false -> None)) funident_map None + in + (match ido with + | Some id -> + ret (Obj.magic coq_Monad_optErr) + (app args' (app (Coq_cons ((Call id), Coq_nil)) t0)) + | None -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Ssequence (s1, s2) -> + bind (Obj.magic coq_Monad_optErr) + (wasm_statement funident_map temps return_type s1 cstack_depth) + (fun x1 -> + bind (Obj.magic coq_Monad_optErr) + (wasm_statement funident_map temps return_type s2 cstack_depth) + (fun x2 -> ret (Obj.magic coq_Monad_optErr) (app x1 x2))) + | Sifthenelse (c, strue, sfalse) -> + bind (Obj.magic coq_Monad_optErr) + (wasm_statement funident_map temps return_type sfalse + (add cstack_depth (S O))) (fun nfalse -> + bind (Obj.magic coq_Monad_optErr) + (wasm_statement funident_map temps return_type strue + (add cstack_depth (S O))) (fun ntrue -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps c Coq_true) + (fun cond -> + ret (Obj.magic coq_Monad_optErr) + (app cond (Coq_cons ((If ((BT_valtype None), ntrue, nfalse)), + Coq_nil)))))) + | Sloop sbody -> + bind (Obj.magic coq_Monad_optErr) + (wasm_statement funident_map temps return_type sbody + (add cstack_depth (S O))) (fun n2 -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Loop ((BT_valtype None), + n2)), Coq_nil))) + | Sbreak -> + (match Nat.eqb cstack_depth O with + | Coq_true -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Coq_false -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Br cstack_depth), + Coq_nil))) + | Sreturn retval -> + bind (Obj.magic coq_Monad_optErr) + (match retval with + | Some i -> wasm_expr temps (Etempvar (i, return_type)) Coq_true + | None -> Success (Coq_cons ((Const_nat O), Coq_nil))) (fun rv' -> + ret (Obj.magic coq_Monad_optErr) (app rv' (Coq_cons (Return, Coq_nil)))) + | Shash (ex1, ex2, exo) -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps ex1 Coq_true) + (fun ex1' -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps ex2 Coq_true) + (fun ex2' -> + bind (Obj.magic coq_Monad_optErr) + (match exo with + | Some ex -> wasm_expr temps ex Coq_true + | None -> ret (Obj.magic coq_Monad_optErr) ex2') (fun exo' -> + ret (Obj.magic coq_Monad_optErr) + (app ex2' + (app (Coq_cons ((Memop (Coq_i32 IOp32.Load)), Coq_nil)) + (app exo' + (app (Coq_cons ((Memop (Coq_i32 IOp32.Load)), Coq_nil)) + (app callExternalSha256TwoParam + (store_len ex1' (S (S (S (S (S (S (S (S O))))))))))))))))) + | Stransfer (a, v) -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps a Coq_true) (fun a' -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps v Coq_true) + (fun v' -> + ret (Obj.magic coq_Monad_optErr) + (app (set_eeiOffset a' Coq_addressOffset (S O)) + (app (set_eeiOffset v' Coq_valueOffset (S O)) + (app (set_eeiOffset Coq_nil Coq_dataOffset (S O)) (Coq_cons + ((Const_nat default_gas_limit), (Coq_cons ((Cvtop (Coq_i32 + (IOp32.Extend_i32 Unsigned))), (Coq_cons ((Const_nat + (get_idx_eei_mem_offset Coq_addressOffset)), (Coq_cons + ((Const_nat (get_idx_eei_mem_offset Coq_valueOffset)), + (Coq_cons ((Const_nat + (get_idx_eei_mem_offset Coq_dataOffset)), (Coq_cons + ((Const_nat O), (Coq_cons ((Call (get_idx_eei Coq_eei_call)), + Coq_nil))))))))))))))))))) + | Scallmethod (a, _, _, v, _, args) -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps a Coq_true) (fun a' -> + bind (Obj.magic coq_Monad_optErr) + (fold_left (fun xs x -> + bind (Obj.magic coq_Monad_optErr) xs (fun xsl -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps x Coq_true) + (fun rst -> ret (Obj.magic coq_Monad_optErr) (app xsl rst)))) + args (ret (Obj.magic coq_Monad_optErr) Coq_nil)) (fun _ -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps v Coq_true) + (fun v' -> + ret (Obj.magic coq_Monad_optErr) + (app (set_eeiOffset a' Coq_addressOffset (S O)) + (app (set_eeiOffset v' Coq_valueOffset (S O)) + (app (set_eeiOffset Coq_nil Coq_dataOffset (S O)) (Coq_cons + ((Const_nat default_gas_limit), (Coq_cons ((Cvtop (Coq_i32 + (IOp32.Extend_i32 Unsigned))), (Coq_cons ((Const_nat + (get_idx_eei_mem_offset Coq_addressOffset)), (Coq_cons + ((Const_nat (get_idx_eei_mem_offset Coq_valueOffset)), + (Coq_cons ((Const_nat + (get_idx_eei_mem_offset Coq_dataOffset)), (Coq_cons + ((Const_nat O), (Coq_cons ((Call + (get_idx_eei Coq_eei_call)), Coq_nil)))))))))))))))))))) + | Slog (topics, args) -> + bind (Obj.magic coq_Monad_optErr) + (fold_left (fun xs x -> + bind (Obj.magic coq_Monad_optErr) xs (fun xsl -> + bind (Obj.magic coq_Monad_optErr) + (match x with + | Econst_int256 (i, _) -> + Success (Coq_cons ((Const_256 + (Int256.shru + (Int256.shl i + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO (Coq_xI (Coq_xI Coq_xH)))))))))) + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 + (Int256.shru + (Int256.shl i + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO (Coq_xO (Coq_xI Coq_xH)))))))))) + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 + (Int256.shru + (Int256.shl i + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO (Coq_xI (Coq_xO Coq_xH)))))))))) + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 + (Int256.shru + (Int256.shl i + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO (Coq_xO (Coq_xO Coq_xH)))))))))) + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 + (Int256.shru + (Int256.shl i + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO (Coq_xI Coq_xH))))))))) + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 + (Int256.shru + (Int256.shl i + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO (Coq_xO Coq_xH))))))))) + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 + (Int256.shru + (Int256.shl i + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO Coq_xH)))))))) + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 + (Int256.shru (Int256.shl i (Int256.repr Z0)) + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xI (Coq_xI Coq_xH))))))))))), + Coq_nil)))))))))))))))) + | _ -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))) + (fun rst -> ret (Obj.magic coq_Monad_optErr) (app xsl rst)))) + topics (ret (Obj.magic coq_Monad_optErr) Coq_nil)) (fun topics' -> + bind (Obj.magic coq_Monad_optErr) + (fold_left (fun xs x -> + bind (Obj.magic coq_Monad_optErr) xs (fun xsl -> + bind (Obj.magic coq_Monad_optErr) (wasm_expr temps x Coq_true) + (fun rst -> ret (Obj.magic coq_Monad_optErr) (app xsl rst)))) + args (ret (Obj.magic coq_Monad_optErr) Coq_nil)) (fun args' -> + let topicsnumber = + match topics with + | Coq_nil -> O + | Coq_cons (_, l) -> + (match l with + | Coq_nil -> S O + | Coq_cons (_, l0) -> + (match l0 with + | Coq_nil -> S (S O) + | Coq_cons (_, l1) -> + (match l1 with + | Coq_nil -> S (S (S O)) + | Coq_cons (_, l2) -> + (match l2 with + | Coq_nil -> S (S (S (S O))) + | Coq_cons (_, _) -> S (S (S (S (S O)))))))) + in + let argsnumber = length args in + let compiledInstrs = + app args' + (app (storeLogArgs argsnumber) + (app topics' + (app + (match topics with + | Coq_nil -> Coq_nil + | Coq_cons (_, l) -> + (match l with + | Coq_nil -> + set_eeiOffset Coq_nil Coq_topic1Offset (S (S (S (S (S + (S (S (S O)))))))) + | Coq_cons (_, l0) -> + (match l0 with + | Coq_nil -> + app + (set_eeiOffset Coq_nil Coq_topic1Offset (S (S (S + (S (S (S (S (S O))))))))) + (set_eeiOffset Coq_nil Coq_topic2Offset (S (S (S + (S (S (S (S (S O))))))))) + | Coq_cons (_, l1) -> + (match l1 with + | Coq_nil -> + app + (set_eeiOffset Coq_nil Coq_topic1Offset (S (S + (S (S (S (S (S (S O))))))))) + (app + (set_eeiOffset Coq_nil Coq_topic2Offset (S + (S (S (S (S (S (S (S O))))))))) + (set_eeiOffset Coq_nil Coq_topic3Offset (S + (S (S (S (S (S (S (S O)))))))))) + | Coq_cons (_, l2) -> + (match l2 with + | Coq_nil -> + app + (set_eeiOffset Coq_nil Coq_topic1Offset (S + (S (S (S (S (S (S (S O))))))))) + (app + (set_eeiOffset Coq_nil Coq_topic2Offset + (S (S (S (S (S (S (S (S O))))))))) + (app + (set_eeiOffset Coq_nil + Coq_topic3Offset (S (S (S (S (S (S + (S (S O))))))))) + (set_eeiOffset Coq_nil + Coq_topic4Offset (S (S (S (S (S (S + (S (S O))))))))))) + | Coq_cons (_, _) -> Coq_nil))))) (Coq_cons + ((Const_nat (get_idx_eei_mem_offset Coq_dataOffset)), + (Coq_cons ((Const_nat (mul argsnumber (S (S (S (S O)))))), + (Coq_cons ((Const_nat topicsnumber), (Coq_cons ((Const_nat + (get_idx_eei_mem_offset Coq_topic1Offset)), (Coq_cons + ((Const_nat (get_idx_eei_mem_offset Coq_topic2Offset)), + (Coq_cons ((Const_nat + (get_idx_eei_mem_offset Coq_topic3Offset)), (Coq_cons + ((Const_nat (get_idx_eei_mem_offset Coq_topic4Offset)), + (Coq_cons ((Call (get_idx_eei Coq_eei_log)), + Coq_nil))))))))))))))))))) + in + (match Nat.eqb topicsnumber (S (S (S (S (S O))))) with + | Coq_true -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))) + | Coq_false -> ret (Obj.magic coq_Monad_optErr) compiledInstrs))) + | Srevert -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Const_nat + (get_idx_eei_mem_offset Coq_dataOffset)), (Coq_cons ((Const_nat + eei_revert_output_data_length), (Coq_cons ((Call + (get_idx_eei Coq_eei_revert)), Coq_nil)))))) + +(** val allocate_temps : (ident, coq_type) prod list -> nat -> nat PTree.t **) + +let rec allocate_temps ids base = + match ids with + | Coq_nil -> PTree.empty + | Coq_cons (p, rest) -> + let Coq_pair (id, _) = p in + PTree.set id base (allocate_temps rest (S base)) + +(** val allocate_all_temps : (ident, coq_type) prod list -> nat PTree.t **) + +let allocate_all_temps ids = + allocate_temps ids O + +(** val allocate_fn_temps : coq_function -> nat PTree.t **) + +let allocate_fn_temps f = + allocate_all_temps (app f.fn_params f.fn_temps) + +(** val wasm_function : + coq_Z PTree.t -> (positive, nat) prod list -> coq_function -> + Language0.coq_function optErr **) + +let wasm_function _ funident_map f = + let wasm_fb = + wasm_statement funident_map (allocate_fn_temps f) f.fn_return f.fn_body O + in + (match wasm_fb with + | Success insts -> + ret (Obj.magic coq_Monad_optErr) { Language0.fn_return = f.fn_return; + Language0.fn_params = f.fn_params; Language0.fn_temps = f.fn_temps; + Language0.fn_locals = f.fn_locals; Language0.fn_body = + (app (Coq_cons ((Const_256 sp), (Coq_cons ((Const_256 sb), (Coq_cons + ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))))) insts) } + | Error msg -> + Error + (append (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + msg)) + +(** val wasm_constructor : + coq_function option -> coq_Z PTree.t -> (positive, nat) prod list -> + Language0.coq_function option optErr **) + +let wasm_constructor f global funident_map = + match f with + | Some f0 -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic wasm_function global funident_map f0) (fun cf -> + ret (Obj.magic coq_Monad_optErr) (Some cf)) + | None -> ret (Obj.magic coq_Monad_optErr) None + +(** val wasm_functions : + coq_function PTree.t -> coq_Z PTree.t -> (positive, nat) prod list -> + Language0.coq_function PTree.t optErr **) + +let wasm_functions defs global funident_map = + transl_tree (wasm_function global funident_map) defs + +(** val wasm_methoddefs : + coq_function option IntMap.t -> coq_Z PTree.t -> (positive, nat) prod + list -> Language0.coq_function option IntMap.t optErr **) + +let wasm_methoddefs defs global funident_map = + transl_map (wasm_function global funident_map) defs + +(** val wasm_func_identifier : + (positive, 'a1) prod list -> nat -> nat -> (positive, nat) prod list **) + +let rec wasm_func_identifier a idx base = + match a with + | Coq_nil -> Coq_nil + | Coq_cons (x, xs) -> + Coq_cons ((Coq_pair ((fst x), (add idx base))), + (wasm_func_identifier xs (add idx (S O)) base)) + +(** val wasm_genv : genv -> Language0.genv optErr **) + +let wasm_genv ge = + let vars = ge.Genv.genv_vars in + let names = ge.Genv.genv_funcs in + let fundefs = ge.Genv.genv_fundefs in + let sigs = ge.Genv.genv_methods in + let defs = ge.Genv.genv_defs in + let methoddefs = ge.Genv.genv_methoddefs in + let constructor = ge.Genv.genv_constructor in + let wasm_func_id_base = + add (length sigs) (match constructor with + | Some _ -> S O + | None -> O) + in + let funident_map = + wasm_func_identifier (PTree.elements fundefs) O + (add wasm_func_id_base (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S + O))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + in + bind (Obj.magic coq_Monad_optErr) + (Obj.magic allocate_addrs ge.Genv.genv_vars Z0 ge.Genv.genv_defs) + (fun allocated -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic wasm_functions fundefs allocated funident_map) + (fun fundefs0 -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic wasm_methoddefs methoddefs allocated funident_map) + (fun methoddefs0 -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic wasm_constructor constructor allocated funident_map) + (fun constructor0 -> + ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; + Genv.genv_funcs = names; Genv.genv_methods = sigs; + Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; + Genv.genv_methoddefs = methoddefs0; Genv.genv_constructor = + constructor0 })))) diff --git a/src/backend/extraction/Gen10.mli b/src/backend/extraction/Gen10.mli new file mode 100644 index 0000000..b9dd566 --- /dev/null +++ b/src/backend/extraction/Gen10.mli @@ -0,0 +1,63 @@ +open AST +open Ascii +open BinNums +open BinPos +open Cop +open Ctypes +open Datatypes +open ExpCintptr +open GlobalenvCompile +open Globalenvs +open Int0 +open Integers +open Language0 +open List0 +open Maps0 +open Monad +open Nat0 +open OptErrMonad +open Options +open PeanoNat +open StackEnv +open StmtCintptr +open String0 +open Structure +open Trees +open Values + +val wasm_expr : nat PTree.t -> ExpCintptr.expr -> bool -> instr list optErr + +val wasm_exprs : nat PTree.t -> ExpCintptr.expr list -> instr list optErr + +val optident : nat PTree.t -> ident option -> instrs optErr + +val wasm_statement : + (positive, nat) prod list -> nat PTree.t -> coq_type -> statement -> nat -> + instrs optErr + +val allocate_temps : (ident, coq_type) prod list -> nat -> nat PTree.t + +val allocate_all_temps : (ident, coq_type) prod list -> nat PTree.t + +val allocate_fn_temps : coq_function -> nat PTree.t + +val wasm_function : + coq_Z PTree.t -> (positive, nat) prod list -> coq_function -> + Language0.coq_function optErr + +val wasm_constructor : + coq_function option -> coq_Z PTree.t -> (positive, nat) prod list -> + Language0.coq_function option optErr + +val wasm_functions : + coq_function PTree.t -> coq_Z PTree.t -> (positive, nat) prod list -> + Language0.coq_function PTree.t optErr + +val wasm_methoddefs : + coq_function option IntMap.t -> coq_Z PTree.t -> (positive, nat) prod list + -> Language0.coq_function option IntMap.t optErr + +val wasm_func_identifier : + (positive, 'a1) prod list -> nat -> nat -> (positive, nat) prod list + +val wasm_genv : genv -> Language0.genv optErr diff --git a/src/backend/extraction/Gen2.ml b/src/backend/extraction/Gen2.ml index e5d0ebe..3e6390b 100644 --- a/src/backend/extraction/Gen2.ml +++ b/src/backend/extraction/Gen2.ml @@ -1,159 +1,1639 @@ +open AST open Ascii +open BinNums +open BinPos open Datatypes +open ExpCintptr open Globalenvs +open List0 open Maps0 open Monad +open Nat0 open OptErrMonad -open Semantics +open PeanoNat open StmtCGraph -open StmtClinear open String0 open Trees -(** val cbasic_stm : - bool option -> node -> StmtCGraph.statement -> bblock **) - -let cbasic_stm ismethod nthis = function -| Sskip n -> Coq_cons ((Sjump n), Coq_nil) -| StmtCGraph.Smassign (lv, rv, n) -> - Coq_cons ((Smassign (lv, rv)), (Coq_cons ((Sjump n), Coq_nil))) -| StmtCGraph.Ssassign (lv, rv, n) -> - Coq_cons ((Ssassign (lv, rv)), (Coq_cons ((Sjump n), Coq_nil))) -| StmtCGraph.Sset (id, rv, n) -> - Coq_cons ((Sset (id, rv)), (Coq_cons ((Sjump n), Coq_nil))) -| StmtCGraph.Scall (retval, lab, args, n) -> - Coq_cons ((Scall (retval, lab, args, nthis)), (Coq_cons ((Sjump n), - Coq_nil))) -| Scond (cond, ntrue, nfalse) -> - Coq_cons ((Sjumpi (cond, ntrue)), (Coq_cons ((Sjump nfalse), Coq_nil))) -| StmtCGraph.Sreturn (var, n) -> - Coq_cons ((Sreturn var), (Coq_cons ((Sjump n), Coq_nil))) -| StmtCGraph.Sdone -> Coq_cons ((Sdone ismethod), Coq_nil) -| StmtCGraph.Shash (ex1, ex2, exo, n) -> - Coq_cons ((Shash (ex1, ex2, exo)), (Coq_cons ((Sjump n), Coq_nil))) -| StmtCGraph.Stransfer (a, v, nfail, n) -> - Coq_cons ((Stransfer (a, v, nfail)), (Coq_cons ((Sjump n), Coq_nil))) -| StmtCGraph.Scallmethod (a, rvs, sig0, v, args, nfail, n) -> - Coq_cons ((Scallmethod (a, rvs, sig0, v, args, nfail)), (Coq_cons ((Sjump - n), Coq_nil))) -| StmtCGraph.Slog (topics, args, n) -> - Coq_cons ((Slog (topics, args)), (Coq_cons ((Sjump n), Coq_nil))) -| StmtCGraph.Srevert -> Coq_cons (Srevert, Coq_nil) - -(** val cbasic_code : bool option -> StmtCGraph.code -> Semantics.code **) - -let cbasic_code ismethod c = - PTree.map (cbasic_stm ismethod) c - -(** val cbasic_function : - bool option -> StmtCGraph.coq_function -> Semantics.coq_function **) - -let cbasic_function ismethod f = - { Semantics.fn_return = f.StmtCGraph.fn_return; Semantics.fn_params = - f.StmtCGraph.fn_params; Semantics.fn_temps = f.StmtCGraph.fn_temps; - Semantics.fn_locals = f.StmtCGraph.fn_locals; Semantics.fn_code = - (cbasic_code ismethod f.StmtCGraph.fn_code); Semantics.fn_entrypoint = - f.fn_entrypoint } - -(** val cbasic_fundef : - bool option -> StmtCGraph.coq_function -> Semantics.coq_function optErr **) - -let cbasic_fundef ismethod f = - ret (Obj.magic coq_Monad_optErr) (cbasic_function ismethod f) - -(** val cbasic_fundefs : - StmtCGraph.coq_function PTree.t -> Semantics.coq_function PTree.t optErr **) - -let cbasic_fundefs t0 = - transl_tree (cbasic_fundef (Some Coq_false)) t0 - -(** val cbasic_methoddefs : - StmtCGraph.coq_function option IntMap.t -> Semantics.coq_function option - IntMap.t optErr **) - -let cbasic_methoddefs methods = - transl_map (cbasic_fundef (Some Coq_true)) methods - -(** val cbasic_constructor : - StmtCGraph.coq_function option -> Semantics.coq_function optErr **) - -let cbasic_constructor = function -| Some c -> - bind (Obj.magic coq_Monad_optErr) (cbasic_fundef None c) (fun f -> - ret (Obj.magic coq_Monad_optErr) f) -| None -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - -(** val cbasic_genv : StmtCGraph.genv -> Semantics.genv optErr **) - -let cbasic_genv ge = +type set = bool PTree.t + +type coq_LVDomain = set + +(** val set_empty : set **) + +let set_empty = + PTree.empty + +(** val set_union' : positive list -> set -> set -> set -> set **) + +let rec set_union' keys x y r = + match keys with + | Coq_nil -> r + | Coq_cons (hd, rest) -> + set_union' rest x y + (PTree.set hd + (match PTree.get_default Coq_false hd x with + | Coq_true -> Coq_true + | Coq_false -> PTree.get_default Coq_false hd y) r) + +(** val set_union : set -> set -> set **) + +let set_union x y = + set_union' (app (PTree.xkeys x Coq_xH) (PTree.xkeys y Coq_xH)) x y + PTree.empty + +(** val set_minus' : positive list -> set -> set -> set -> set **) + +let rec set_minus' keys x y r = + match keys with + | Coq_nil -> r + | Coq_cons (hd, rest) -> + set_minus' rest x y + (PTree.set hd + (match PTree.get_default Coq_false hd x with + | Coq_true -> negb (PTree.get_default Coq_false hd y) + | Coq_false -> Coq_false) r) + +(** val set_minus : bool PTree.t -> set -> set **) + +let set_minus x y = + set_minus' (PTree.xkeys x Coq_xH) x y PTree.empty + +(** val set_add : ident -> set -> set **) + +let rec set_add i r = + PTree.set i Coq_true r + +(** val set_in : ident -> set -> bool **) + +let rec set_in i r = + PTree.get_default Coq_false i r + +(** val set_card : set -> nat **) + +let rec set_card x = + PTree.fold1 (fun acc e -> + match e with + | Coq_true -> add acc (S O) + | Coq_false -> acc) x O + +(** val set_eq : set -> set -> bool **) + +let rec set_eq x y = + match Nat.eqb (set_card (set_minus x y)) O with + | Coq_true -> Nat.eqb (set_card (set_minus y x)) O + | Coq_false -> Coq_false + +(** val set_fold_left : ('a1 -> ident -> 'a1) -> set -> 'a1 -> 'a1 **) + +let set_fold_left f m v = + PTree.fold (fun acc i e -> + match e with + | Coq_true -> f acc i + | Coq_false -> acc) m v + +(** val get_use_expr : expr -> coq_LVDomain **) + +let rec get_use_expr = function +| Etempvar (i, _) -> set_add i set_empty +| Esload (e0, _) -> get_use_expr e0 +| Emload (e0, _) -> get_use_expr e0 +| Eaddr (e0, _) -> get_use_expr e0 +| Eunop (_, e0, _) -> get_use_expr e0 +| Ebinop (_, e1, e2, _) -> set_union (get_use_expr e1) (get_use_expr e2) +| Ecall1 (_, e0, _) -> get_use_expr e0 +| _ -> set_empty + +(** val get_use_exprs : expr list -> coq_LVDomain **) + +let rec get_use_exprs = function +| Coq_nil -> set_empty +| Coq_cons (e, rest) -> set_union (get_use_expr e) (get_use_exprs rest) + +(** val get_use : statement -> coq_LVDomain **) + +let get_use = function +| Smassign (lv, rv, _) -> set_union (get_use_expr lv) (get_use_expr rv) +| Ssassign (lv, rv, _) -> set_union (get_use_expr lv) (get_use_expr rv) +| Sset (_, rv, _) -> get_use_expr rv +| Scall (_, _, args, _) -> get_use_exprs args +| Scond (cond, _, _) -> get_use_expr cond +| Sreturn (val0, _) -> + (match val0 with + | Some e -> set_add e set_empty + | None -> set_empty) +| Shash (e1, e2, eo, _) -> + let e1' = get_use_expr e1 in + let e2' = get_use_expr e2 in + (match eo with + | Some eoo -> + let eoo' = get_use_expr eoo in set_union (set_union e1' e2') eoo' + | None -> set_union e1' e2') +| Stransfer (a, v, _, _) -> set_union (get_use_expr a) (get_use_expr v) +| Scallmethod (a, _, _, v, _, args, _, _) -> + set_union (set_union (get_use_expr a) (get_use_expr v)) (get_use_exprs args) +| Slog (topics, args, _) -> + set_union (get_use_exprs topics) (get_use_exprs args) +| _ -> set_empty + +(** val get_successor : statement -> node list **) + +let get_successor = function +| Sskip n -> Coq_cons (n, Coq_nil) +| Smassign (_, _, n) -> Coq_cons (n, Coq_nil) +| Ssassign (_, _, n) -> Coq_cons (n, Coq_nil) +| Sset (_, _, n) -> Coq_cons (n, Coq_nil) +| Scall (_, _, _, n) -> Coq_cons (n, Coq_nil) +| Scond (_, ntrue, nfalse) -> Coq_cons (ntrue, (Coq_cons (nfalse, Coq_nil))) +| Sreturn (_, n) -> Coq_cons (n, Coq_nil) +| Shash (_, _, _, n) -> Coq_cons (n, Coq_nil) +| Stransfer (_, _, nfail, n) -> Coq_cons (nfail, (Coq_cons (n, Coq_nil))) +| Scallmethod (_, _, _, _, _, _, nfail, n) -> + Coq_cons (nfail, (Coq_cons (n, Coq_nil))) +| Slog (_, _, n) -> Coq_cons (n, Coq_nil) +| _ -> Coq_nil + +(** val get_def : statement -> coq_LVDomain **) + +let get_def = function +| Sset (id, _, _) -> set_add id set_empty +| Scall (retval, _, _, _) -> + (match retval with + | Some rv -> set_add rv set_empty + | None -> set_empty) +| Scallmethod (_, rvs, _, _, _, _, _, _) -> + fold_left (fun acc e -> set_add e acc) rvs set_empty +| _ -> set_empty + +type coq_CD = coq_LVDomain PTree.t + +(** val f_b : code -> node -> coq_CD -> coq_LVDomain **) + +let f_b cfg block cd = + match PTree.get block cfg with + | Some s -> + let oUT = + fold_left (fun acc e -> + set_union acc (PTree.get_default set_empty e cd)) (get_successor s) + set_empty + in + set_union (get_use s) (set_minus oUT (get_def s)) + | None -> set_empty + +(** val update_at : code -> coq_CD -> node -> coq_CD **) + +let update_at cfg cd b = + PTree.set b (f_b cfg b cd) cd + +(** val lv_top : code -> coq_LVDomain **) + +let lv_top cfg = + PTree.fold (fun s _ stmt -> set_union s (get_use stmt)) cfg set_empty + +(** val cd_top : code -> coq_CD **) + +let cd_top cfg = + PTree.map1 (fun _ -> lv_top cfg) cfg + +(** val get_lv_once : code -> node list -> coq_CD -> coq_CD **) + +let get_lv_once cfg nodes cd = + fold_left (update_at cfg) nodes cd + +(** val get_lv_mfp : code -> node list -> nat -> coq_CD -> coq_CD **) + +let rec get_lv_mfp cfg blocks iteration cd = + match iteration with + | O -> cd + | S x -> + let new_cd = get_lv_once cfg blocks cd in + (match PTree.beq set_eq cd new_cd with + | Coq_true -> new_cd + | Coq_false -> get_lv_mfp cfg blocks x new_cd) + +(** val compute_livevar : code -> nat -> coq_CD optErr **) + +let compute_livevar cfg precision = + let blocks = map fst (PTree.elements cfg) in + ret (Obj.magic coq_Monad_optErr) + (get_lv_mfp cfg blocks precision (cd_top cfg)) + +type tvs = set + +type clashg = tvs PTree.t + +(** val create_clash_graph : coq_CD -> clashg optErr **) + +let rec create_clash_graph lv = + ret (Obj.magic coq_Monad_optErr) + (PTree.fold1 (fun acc e -> + set_fold_left (fun acc' e' -> + PTree.set e' (set_union (PTree.get_default set_empty e' acc') e) acc') + e acc) lv PTree.empty) + +(** val remove_node_graph : ident -> clashg -> clashg **) + +let rec remove_node_graph t0 g = + let g' = PTree.remove t0 g in + PTree.fold (fun acc i e -> + PTree.set i (set_minus e (set_add t0 set_empty)) acc) g' g' + +(** val least_degree : clashg -> ident **) + +let rec least_degree g = + let Coq_pair (id, _) = + PTree.fold (fun acc i e -> + let Coq_pair (_, mn) = acc in + let c = set_card e in + (match Nat.leb c mn with + | Coq_true -> Coq_pair (i, c) + | Coq_false -> acc)) g (Coq_pair ((Pos.of_nat O), (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S O)))))))))))))))))) + in + id + +(** val pop_graph : nat -> clashg -> ident list -> ident list optErr **) + +let rec pop_graph i g stack = + match i with + | O -> + (match Nat.eqb (length (PTree.elements g)) O with + | Coq_true -> ret (Obj.magic coq_Monad_optErr) stack + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | S n -> + let ldn = least_degree g in + pop_graph n (remove_node_graph ldn g) (Coq_cons (ldn, stack)) + +type colorMap = nat PTree.t + +(** val get_assigned_colors : colorMap -> tvs -> set **) + +let get_assigned_colors c s = + PTree.fold (fun acc i e -> + match e with + | Coq_true -> + (match PTree.get i c with + | Some cl -> set_add (Pos.of_nat cl) acc + | None -> acc) + | Coq_false -> acc) s set_empty + +(** val colorMapFull : set **) + +let colorMapFull = + set_add (Pos.of_nat (S O)) + (set_add (Pos.of_nat (S (S O))) + (set_add (Pos.of_nat (S (S (S O)))) + (set_add (Pos.of_nat (S (S (S (S O))))) + (set_add (Pos.of_nat (S (S (S (S (S O)))))) + (set_add (Pos.of_nat (S (S (S (S (S (S O))))))) + (set_add (Pos.of_nat (S (S (S (S (S (S (S O)))))))) + (set_add (Pos.of_nat (S (S (S (S (S (S (S (S O))))))))) + (set_add (Pos.of_nat (S (S (S (S (S (S (S (S (S O)))))))))) + (set_add + (Pos.of_nat (S (S (S (S (S (S (S (S (S (S O))))))))))) + (set_add + (Pos.of_nat (S (S (S (S (S (S (S (S (S (S (S + O)))))))))))) + (set_add + (Pos.of_nat (S (S (S (S (S (S (S (S (S (S (S (S + O))))))))))))) + (set_add + (Pos.of_nat (S (S (S (S (S (S (S (S (S (S (S (S + (S O)))))))))))))) + (set_add + (Pos.of_nat (S (S (S (S (S (S (S (S (S (S (S (S + (S (S O))))))))))))))) + (set_add + (Pos.of_nat (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S O)))))))))))))))) set_empty)))))))))))))) + +(** val max_reg : nat **) + +let max_reg = + S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S O))))))))))))))) + +(** val default_reg : nat **) + +let default_reg = + S O + +(** val assign_color : ident list -> clashg -> colorMap -> colorMap optErr **) + +let rec assign_color stack g c = + match stack with + | Coq_nil -> ret (Obj.magic coq_Monad_optErr) c + | Coq_cons (x, xs) -> + (match PTree.get x g with + | Some lk -> + let assignable = set_minus colorMapFull (get_assigned_colors c lk) in + (match Nat.ltb O (set_card assignable) with + | Coq_true -> + (match PTree.fold (fun acc i e -> + match e with + | Coq_true -> Some i + | Coq_false -> acc) assignable None with + | Some reg -> assign_color xs g (PTree.set x (Pos.to_nat reg) c) + | None -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))) + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | None -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + +(** val recolor_expr : expr -> colorMap -> expr optErr **) + +let rec recolor_expr e c = + match e with + | Etempvar (i, t0) -> + ret (Obj.magic coq_Monad_optErr) (Etempvar + ((Pos.of_nat (PTree.get_default default_reg i c)), t0)) + | Esload (e0, t0) -> + bind (Obj.magic coq_Monad_optErr) (recolor_expr e0 c) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (Esload (e', t0))) + | Emload (e0, t0) -> + bind (Obj.magic coq_Monad_optErr) (recolor_expr e0 c) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (Emload (e', t0))) + | Eaddr (e0, t0) -> + bind (Obj.magic coq_Monad_optErr) (recolor_expr e0 c) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (Eaddr (e', t0))) + | Eunop (o, e0, t0) -> + bind (Obj.magic coq_Monad_optErr) (recolor_expr e0 c) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (Eunop (o, e', t0))) + | Ebinop (o, e1, e2, t0) -> + bind (Obj.magic coq_Monad_optErr) (recolor_expr e1 c) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) (recolor_expr e2 c) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) (Ebinop (o, e1', e2', t0)))) + | Ecall1 (b, e0, t0) -> + bind (Obj.magic coq_Monad_optErr) (recolor_expr e0 c) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (Ecall1 (b, e', t0))) + | _ -> ret (Obj.magic coq_Monad_optErr) e + +(** val recolor_exprs : expr list -> colorMap -> expr list optErr **) + +let rec recolor_exprs es c = + match es with + | Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil + | Coq_cons (e, rest) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr e c) (fun e' -> + bind (Obj.magic coq_Monad_optErr) (recolor_exprs rest c) (fun rest' -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (e', rest')))) + +(** val recolor_regs : + ident list -> colorMap -> ident list -> ident list optErr **) + +let rec recolor_regs l c r = + match l with + | Coq_nil -> ret (Obj.magic coq_Monad_optErr) r + | Coq_cons (x, xs) -> + recolor_regs xs c (Coq_cons + ((Pos.of_nat (PTree.get_default default_reg x c)), r)) + +(** val recolor_reg : ident -> colorMap -> ident **) + +let recolor_reg i c = + Pos.of_nat (PTree.get_default default_reg i c) + +(** val recolor_expr_opt : expr option -> colorMap -> expr option optErr **) + +let recolor_expr_opt e c = + match e with + | Some e' -> + (match recolor_expr e' c with + | Success s -> Success (Some s) + | Error s -> Error s) + | None -> Success None + +(** val recolor : + code -> set -> nat -> node -> colorMap -> (code -> code) -> (code -> + code) optErr **) + +let rec recolor cfg t0 rd tn c cont = + match rd with + | O -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | S x -> + (match PTree.get tn cfg with + | Some s -> + (match set_in tn t0 with + | Coq_true -> ret (Obj.magic coq_Monad_optErr) cont + | Coq_false -> + let t' = set_union t0 (set_add tn set_empty) in + let recolor' = recolor cfg t' x in + (match s with + | Sskip n -> recolor' n c cont + | Smassign (lv, rv, n) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr lv c) + (fun lv' -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_expr rv c) (fun rv' -> + recolor' n c (fun g -> + PTree.set tn (Smassign (lv', rv', n)) (cont g)))) + | Ssassign (lv, rv, n) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr lv c) + (fun lv' -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_expr rv c) (fun rv' -> + recolor' n c (fun g -> + PTree.set tn (Ssassign (lv', rv', n)) (cont g)))) + | Sset (id, rv, n) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr rv c) + (fun rv' -> + recolor' n c (fun g -> + PTree.set tn (Sset ((recolor_reg id c), rv', n)) (cont g))) + | Scall (retval, lab, args, n) -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_exprs args c) (fun args' -> + match retval with + | Some rv -> + recolor' n c (fun g -> + PTree.set tn (Scall ((Some (recolor_reg rv c)), lab, + args', n)) (cont g)) + | None -> + recolor' n c (fun g -> + PTree.set tn (Scall (retval, lab, args', n)) (cont g))) + | Scond (cond, ntrue, nfalse) -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_expr cond c) (fun cond' -> + bind (Obj.magic coq_Monad_optErr) + (recolor' ntrue c (fun g -> g)) (fun cont' -> + recolor' nfalse c (fun g -> + PTree.set tn (Scond (cond', ntrue, nfalse)) + (cont' (cont g))))) + | Sreturn (val0, n) -> + (match val0 with + | Some v -> + recolor' n c (fun g -> + PTree.set tn (Sreturn ((Some (recolor_reg v c)), n)) + (cont g)) + | None -> recolor' n c cont) + | Shash (e1, e2, eo, n) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr e1 c) + (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_expr e2 c) (fun e2' -> + match eo with + | Some eoo -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_expr eoo c) (fun eoo' -> + recolor' n c (fun g -> + PTree.set tn (Shash (e1', e2', (Some eoo'), n)) + (cont g))) + | None -> + recolor' n c (fun g -> + PTree.set tn (Shash (e1', e2', None, n)) (cont g)))) + | Stransfer (a, v, nfail, n) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr a c) + (fun a' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr v c) + (fun v' -> + bind (Obj.magic coq_Monad_optErr) + (recolor' nfail c (fun g -> g)) (fun cont' -> + recolor' n c (fun g -> + PTree.set tn (Stransfer (a', v', nfail, n)) + (cont' (cont g)))))) + | Scallmethod (a, rvs, sig0, v, gas, args, nfail, n) -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_regs rvs c Coq_nil) (fun rvs' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic recolor_expr a c) + (fun a' -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_expr v c) (fun v' -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_expr_opt gas c) (fun gas' -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_exprs args c) (fun args' -> + bind (Obj.magic coq_Monad_optErr) + (recolor' nfail c (fun g -> g)) (fun cont' -> + recolor' n c (fun g -> + PTree.set tn (Scallmethod (a', rvs', sig0, v', + gas', args', nfail, n)) (cont' (cont g))))))))) + | Slog (topics, args, n) -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_exprs topics c) (fun topics' -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor_exprs args c) (fun args' -> + recolor' n c (fun g -> + PTree.set tn (Slog (topics', args', n)) (cont g)))) + | _ -> ret (Obj.magic coq_Monad_optErr) cont)) + | None -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + +(** val color_graph : + code -> nat -> node -> clashg -> (code, colorMap) prod optErr **) + +let color_graph cfg maxnode nentry cg = + match PTree.fold (fun acc _ e -> + match acc with + | Coq_true -> Nat.ltb (set_card e) max_reg + | Coq_false -> Coq_false) cg Coq_true with + | Coq_true -> + let num_tv = length (PTree.elements cg) in + bind (Obj.magic coq_Monad_optErr) + (Obj.magic pop_graph (length (PTree.elements cg)) cg Coq_nil) + (fun stack -> + let unique_nodes = + Nat.eqb + (set_card (fold_left (fun acc e -> set_add e acc) stack set_empty)) + num_tv + in + (match unique_nodes with + | Coq_true -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic assign_color stack cg PTree.empty) (fun colors -> + match Nat.eqb (length (PTree.elements colors)) num_tv with + | Coq_true -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic recolor cfg set_empty maxnode nentry colors + (fun g -> g)) (fun f -> + ret (Obj.magic coq_Monad_optErr) (Coq_pair ((f cfg), colors))) + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Coq_false -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + +(** val variable_coalescing : + code -> nat -> node -> nat -> (code, colorMap) prod optErr **) + +let variable_coalescing cfg maxnode nentry precision = + bind (Obj.magic coq_Monad_optErr) (Obj.magic compute_livevar cfg precision) + (fun lvs -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic create_clash_graph lvs) + (fun cg -> color_graph cfg maxnode nentry cg)) + +(** val get_clash_graph : + code -> nat -> (positive, positive list) prod list optErr **) + +let get_clash_graph cfg precision = + bind (Obj.magic coq_Monad_optErr) (Obj.magic compute_livevar cfg precision) + (fun lvs -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic create_clash_graph lvs) + (fun cg -> + ret (Obj.magic coq_Monad_optErr) + (map (fun x -> + let Coq_pair (id, tvss) = x in + Coq_pair (id, + (PTree.fold (fun acc i ind -> + match ind with + | Coq_true -> Coq_cons (i, acc) + | Coq_false -> acc) tvss Coq_nil))) (PTree.elements cg)))) + +(** val max_iteration : nat **) + +let max_iteration = + S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S + (S (S (S (S (S (S (S (S (S (S + O)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + +(** val copt_function : coq_function -> coq_function optErr **) + +let copt_function f = + match variable_coalescing f.fn_code (sub (Pos.to_nat f.fn_nextnode) (S O)) + f.fn_entrypoint max_iteration with + | Success p -> + let Coq_pair (ncfg, cmap) = p in + ret (Obj.magic coq_Monad_optErr) { fn_return = f.fn_return; fn_params = + (map (fun e -> + let Coq_pair (id, ty) = e in + Coq_pair ((Pos.of_nat (PTree.get_default (S O) id cmap)), ty)) + f.fn_params); fn_temps = + (map (fun e -> + let Coq_pair (id, ty) = e in + Coq_pair ((Pos.of_nat (PTree.get_default (S O) id cmap)), ty)) + f.fn_temps); fn_locals = f.fn_locals; fn_code = ncfg; fn_entrypoint = + f.fn_entrypoint; fn_nextnode = f.fn_nextnode } + | Error msg -> Error msg + +(** val copt_constructor : + coq_function option -> coq_function option optErr **) + +let copt_constructor f = match f with +| Some f0 -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic copt_function f0) (fun cf -> + ret (Obj.magic coq_Monad_optErr) (Some cf)) +| None -> ret (Obj.magic coq_Monad_optErr) f + +(** val clash_function : + coq_function -> (positive, positive list) prod list optErr **) + +let clash_function f = + get_clash_graph f.fn_code max_iteration + +(** val copt_functions : + coq_function PTree.t -> coq_function PTree.t optErr **) + +let copt_functions defs = + transl_tree copt_function defs + +(** val copt_methoddefs : + coq_function option IntMap.t -> coq_function option IntMap.t optErr **) + +let copt_methoddefs defs = + transl_map copt_function defs + +(** val graphviz_helper : + coq_function -> ((positive, statement) prod list, positive) prod **) + +let graphviz_helper f = + Coq_pair ((PTree.elements f.fn_code), f.fn_entrypoint) + +(** val clash_viz : + genv -> (positive, positive list) prod list list optErr **) + +let clash_viz ge = + fold_left (fun acc x -> + match acc with + | Success acc' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic clash_function x) + (fun x' -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (x', acc'))) + | Error msg -> Error msg) (Genv.all_functions ge) (Success Coq_nil) + +(** val cgraph_viz : + genv -> ((positive, statement) prod list, positive) prod list optErr **) + +let cgraph_viz ge = + fold_left (fun acc x -> + match acc with + | Success acc' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic copt_function x) + (fun x' -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((graphviz_helper x'), + acc'))) + | Error msg -> Error msg) (Genv.all_functions ge) (Success Coq_nil) + +(** val copt_genv : genv -> genv optErr **) + +let copt_genv ge = let vars = ge.Genv.genv_vars in - let funcs = ge.Genv.genv_funcs in - let methods = ge.Genv.genv_methods in - let defs = ge.Genv.genv_defs in + let names = ge.Genv.genv_funcs in let fundefs = ge.Genv.genv_fundefs in + let sigs = ge.Genv.genv_methods in + let defs = ge.Genv.genv_defs in let methoddefs = ge.Genv.genv_methoddefs in let constructor = ge.Genv.genv_constructor in - bind (Obj.magic coq_Monad_optErr) (Obj.magic cbasic_fundefs fundefs) + bind (Obj.magic coq_Monad_optErr) (Obj.magic copt_functions fundefs) (fun fundefs0 -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic cbasic_methoddefs methoddefs) (fun methoddefs0 -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic copt_methoddefs methoddefs) + (fun methoddefs0 -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic cbasic_constructor constructor) (fun constructor0 -> + (Obj.magic copt_constructor constructor) (fun constructor0 -> ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; - Genv.genv_funcs = funcs; Genv.genv_methods = methods; - Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; - Genv.genv_methoddefs = methoddefs0; Genv.genv_constructor = (Some - constructor0) }))) + Genv.genv_funcs = names; Genv.genv_methods = sigs; Genv.genv_defs = + defs; Genv.genv_fundefs = fundefs0; Genv.genv_methoddefs = + methoddefs0; Genv.genv_constructor = constructor0 }))) diff --git a/src/backend/extraction/Gen2.mli b/src/backend/extraction/Gen2.mli index b30c22d..dfced14 100644 --- a/src/backend/extraction/Gen2.mli +++ b/src/backend/extraction/Gen2.mli @@ -1,33 +1,137 @@ +open AST open Ascii +open BinNums +open BinPos open Datatypes +open ExpCintptr open Globalenvs +open List0 open Maps0 open Monad +open Nat0 open OptErrMonad -open Semantics +open PeanoNat open StmtCGraph -open StmtClinear open String0 open Trees -val cbasic_stm : bool option -> node -> StmtCGraph.statement -> bblock +type set = bool PTree.t -val cbasic_code : bool option -> StmtCGraph.code -> Semantics.code +type coq_LVDomain = set -val cbasic_function : - bool option -> StmtCGraph.coq_function -> Semantics.coq_function +val set_empty : set -val cbasic_fundef : - bool option -> StmtCGraph.coq_function -> Semantics.coq_function optErr +val set_union' : positive list -> set -> set -> set -> set -val cbasic_fundefs : - StmtCGraph.coq_function PTree.t -> Semantics.coq_function PTree.t optErr +val set_union : set -> set -> set -val cbasic_methoddefs : - StmtCGraph.coq_function option IntMap.t -> Semantics.coq_function option - IntMap.t optErr +val set_minus' : positive list -> set -> set -> set -> set -val cbasic_constructor : - StmtCGraph.coq_function option -> Semantics.coq_function optErr +val set_minus : bool PTree.t -> set -> set -val cbasic_genv : StmtCGraph.genv -> Semantics.genv optErr +val set_add : ident -> set -> set + +val set_in : ident -> set -> bool + +val set_card : set -> nat + +val set_eq : set -> set -> bool + +val set_fold_left : ('a1 -> ident -> 'a1) -> set -> 'a1 -> 'a1 + +val get_use_expr : expr -> coq_LVDomain + +val get_use_exprs : expr list -> coq_LVDomain + +val get_use : statement -> coq_LVDomain + +val get_successor : statement -> node list + +val get_def : statement -> coq_LVDomain + +type coq_CD = coq_LVDomain PTree.t + +val f_b : code -> node -> coq_CD -> coq_LVDomain + +val update_at : code -> coq_CD -> node -> coq_CD + +val lv_top : code -> coq_LVDomain + +val cd_top : code -> coq_CD + +val get_lv_once : code -> node list -> coq_CD -> coq_CD + +val get_lv_mfp : code -> node list -> nat -> coq_CD -> coq_CD + +val compute_livevar : code -> nat -> coq_CD optErr + +type tvs = set + +type clashg = tvs PTree.t + +val create_clash_graph : coq_CD -> clashg optErr + +val remove_node_graph : ident -> clashg -> clashg + +val least_degree : clashg -> ident + +val pop_graph : nat -> clashg -> ident list -> ident list optErr + +type colorMap = nat PTree.t + +val get_assigned_colors : colorMap -> tvs -> set + +val colorMapFull : set + +val max_reg : nat + +val default_reg : nat + +val assign_color : ident list -> clashg -> colorMap -> colorMap optErr + +val recolor_expr : expr -> colorMap -> expr optErr + +val recolor_exprs : expr list -> colorMap -> expr list optErr + +val recolor_regs : ident list -> colorMap -> ident list -> ident list optErr + +val recolor_reg : ident -> colorMap -> ident + +val recolor_expr_opt : expr option -> colorMap -> expr option optErr + +val recolor : + code -> set -> nat -> node -> colorMap -> (code -> code) -> (code -> code) + optErr + +val color_graph : + code -> nat -> node -> clashg -> (code, colorMap) prod optErr + +val variable_coalescing : + code -> nat -> node -> nat -> (code, colorMap) prod optErr + +val get_clash_graph : + code -> nat -> (positive, positive list) prod list optErr + +val max_iteration : nat + +val copt_function : coq_function -> coq_function optErr + +val copt_constructor : coq_function option -> coq_function option optErr + +val clash_function : + coq_function -> (positive, positive list) prod list optErr + +val copt_functions : coq_function PTree.t -> coq_function PTree.t optErr + +val copt_methoddefs : + coq_function option IntMap.t -> coq_function option IntMap.t optErr + +val graphviz_helper : + coq_function -> ((positive, statement) prod list, positive) prod + +val clash_viz : genv -> (positive, positive list) prod list list optErr + +val cgraph_viz : + genv -> ((positive, statement) prod list, positive) prod list optErr + +val copt_genv : genv -> genv optErr diff --git a/src/backend/extraction/Gen3.ml b/src/backend/extraction/Gen3.ml index 93b6b7b..eda2b39 100644 --- a/src/backend/extraction/Gen3.ml +++ b/src/backend/extraction/Gen3.ml @@ -1,157 +1,109 @@ -open AST open Ascii -open BinNums -open Coqlib open Datatypes open Globalenvs open Maps0 open Monad open OptErrMonad open Semantics -open Specif +open StmtCGraph open StmtClinear open String0 open Trees -(** val enumerate_rest : Semantics.code -> node list **) - -let enumerate_rest c = - PTree.xkeys c Coq_xH - -(** val labels_of_bblock : bblock -> label list **) - -let rec labels_of_bblock = function -| Coq_nil -> Coq_nil -| Coq_cons (s, b') -> - (match s with - | Sjump n -> Coq_cons (n, (labels_of_bblock b')) - | Sjumpi (_, n) -> Coq_cons (n, (labels_of_bblock b')) - | _ -> labels_of_bblock b') - -(** val enumerate'_func : (node list, Semantics.code) sigT -> node list **) - -let rec enumerate'_func x = - let todo = projT1 x in - let c = projT2 x in - let enumerate'0 = fun todo0 c0 -> enumerate'_func (Coq_existT (todo0, c0)) - in - (match todo with - | Coq_nil -> enumerate_rest c - | Coq_cons (n, ns) -> - let filtered_var = PTree.get n c in - (match filtered_var with - | Some b -> - Coq_cons (n, - (enumerate'0 (app (labels_of_bblock b) ns) (PTree.remove n c))) - | None -> enumerate'0 ns c)) - -(** val enumerate' : node list -> Semantics.code -> node list **) - -let enumerate' todo c = - enumerate'_func (Coq_existT (todo, c)) - -(** val enumerate : Semantics.coq_function -> node list **) - -let enumerate f = - enumerate' (Coq_cons (f.fn_entrypoint, Coq_nil)) f.Semantics.fn_code - -(** val starts_with : label -> code -> bool **) - -let starts_with lbl = function -| Coq_nil -> Coq_false -| Coq_cons (s, _) -> - (match s with - | Slabel lbl' -> proj_sumbool (peq lbl lbl') - | _ -> Coq_false) - -(** val add_branch : label -> code -> code **) - -let add_branch s k = - match starts_with s k with - | Coq_true -> k - | Coq_false -> Coq_cons ((Sjump s), k) - -(** val clinear_block : bblock -> code -> code **) - -let rec clinear_block b k = - match b with - | Coq_nil -> k - | Coq_cons (i, b') -> - (match i with - | Slabel _ -> clinear_block b' k - | Sjump n -> add_branch n k - | _ -> Coq_cons (i, (clinear_block b' k))) - -(** val clinear_node : Semantics.coq_function -> node -> code -> code **) - -let clinear_node f pc k = - match PTree.get pc f.Semantics.fn_code with - | Some b -> Coq_cons ((Slabel pc), (clinear_block b k)) - | None -> k - -(** val clinear_body : Semantics.coq_function -> node list -> code **) - -let clinear_body f enum = - list_fold_right (clinear_node f) enum Coq_nil - -(** val clinear_function : Semantics.coq_function -> coq_function optErr **) - -let clinear_function f = - let enum = enumerate f in - ret (Obj.magic coq_Monad_optErr) { fn_return = f.Semantics.fn_return; - fn_params = f.Semantics.fn_params; fn_temps = f.Semantics.fn_temps; - fn_locals = f.Semantics.fn_locals; fn_code = - (add_branch f.fn_entrypoint (clinear_body f enum)) } - -(** val clinear_fundef : Semantics.coq_function -> coq_function optErr **) - -let clinear_fundef = - clinear_function - -(** val clinear_fundefs : - Semantics.coq_function PTree.t -> coq_function PTree.t optErr **) - -let clinear_fundefs t0 = - transl_tree clinear_fundef t0 - -(** val clinear_methoddefs : - Semantics.coq_function option IntMap.t -> coq_function option IntMap.t - optErr **) - -let clinear_methoddefs methods = - transl_map clinear_fundef methods - -(** val clinear_constructor : - Semantics.coq_function option -> coq_function optErr **) - -let clinear_constructor = function +(** val cbasic_stm : + bool option -> node -> StmtCGraph.statement -> bblock **) + +let cbasic_stm ismethod nthis = function +| Sskip n -> Coq_cons ((Sjump n), Coq_nil) +| StmtCGraph.Smassign (lv, rv, n) -> + Coq_cons ((Smassign (lv, rv)), (Coq_cons ((Sjump n), Coq_nil))) +| StmtCGraph.Ssassign (lv, rv, n) -> + Coq_cons ((Ssassign (lv, rv)), (Coq_cons ((Sjump n), Coq_nil))) +| StmtCGraph.Sset (id, rv, n) -> + Coq_cons ((Sset (id, rv)), (Coq_cons ((Sjump n), Coq_nil))) +| StmtCGraph.Scall (retval, lab, args, n) -> + Coq_cons ((Scall (retval, lab, args, nthis)), (Coq_cons ((Sjump n), + Coq_nil))) +| Scond (cond, ntrue, nfalse) -> + Coq_cons ((Sjumpi (cond, ntrue)), (Coq_cons ((Sjump nfalse), Coq_nil))) +| StmtCGraph.Sreturn (var, n) -> + Coq_cons ((Sreturn var), (Coq_cons ((Sjump n), Coq_nil))) +| StmtCGraph.Sdone -> Coq_cons ((Sdone ismethod), Coq_nil) +| StmtCGraph.Shash (ex1, ex2, exo, n) -> + Coq_cons ((Shash (ex1, ex2, exo)), (Coq_cons ((Sjump n), Coq_nil))) +| StmtCGraph.Stransfer (a, v, nfail, n) -> + Coq_cons ((Stransfer (a, v, nfail)), (Coq_cons ((Sjump n), Coq_nil))) +| StmtCGraph.Scallmethod (a, rvs, sig0, v, gas, args, nfail, n) -> + Coq_cons ((Scallmethod (a, rvs, sig0, v, gas, args, nfail)), (Coq_cons + ((Sjump n), Coq_nil))) +| StmtCGraph.Slog (topics, args, n) -> + Coq_cons ((Slog (topics, args)), (Coq_cons ((Sjump n), Coq_nil))) +| StmtCGraph.Srevert -> Coq_cons (Srevert, Coq_nil) + +(** val cbasic_code : bool option -> StmtCGraph.code -> Semantics.code **) + +let cbasic_code ismethod c = + PTree.map (cbasic_stm ismethod) c + +(** val cbasic_function : + bool option -> StmtCGraph.coq_function -> Semantics.coq_function **) + +let cbasic_function ismethod f = + { Semantics.fn_return = f.StmtCGraph.fn_return; Semantics.fn_params = + f.StmtCGraph.fn_params; Semantics.fn_temps = f.StmtCGraph.fn_temps; + Semantics.fn_locals = f.StmtCGraph.fn_locals; Semantics.fn_code = + (cbasic_code ismethod f.StmtCGraph.fn_code); Semantics.fn_entrypoint = + f.fn_entrypoint } + +(** val cbasic_fundef : + bool option -> StmtCGraph.coq_function -> Semantics.coq_function optErr **) + +let cbasic_fundef ismethod f = + ret (Obj.magic coq_Monad_optErr) (cbasic_function ismethod f) + +(** val cbasic_fundefs : + StmtCGraph.coq_function PTree.t -> Semantics.coq_function PTree.t optErr **) + +let cbasic_fundefs t0 = + transl_tree (cbasic_fundef (Some Coq_false)) t0 + +(** val cbasic_methoddefs : + StmtCGraph.coq_function option IntMap.t -> Semantics.coq_function option + IntMap.t optErr **) + +let cbasic_methoddefs methods = + transl_map (cbasic_fundef (Some Coq_true)) methods + +(** val cbasic_constructor : + StmtCGraph.coq_function option -> Semantics.coq_function optErr **) + +let cbasic_constructor = function | Some c -> - bind (Obj.magic coq_Monad_optErr) (clinear_fundef c) (fun f -> + bind (Obj.magic coq_Monad_optErr) (cbasic_fundef None c) (fun f -> ret (Obj.magic coq_Monad_optErr) f) | None -> Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), @@ -182,11 +134,11 @@ let clinear_constructor = function Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -(** val clinear_genv : Semantics.genv -> genv optErr **) +(** val cbasic_genv : StmtCGraph.genv -> Semantics.genv optErr **) -let clinear_genv ge = +let cbasic_genv ge = let vars = ge.Genv.genv_vars in let funcs = ge.Genv.genv_funcs in let methods = ge.Genv.genv_methods in @@ -194,12 +146,12 @@ let clinear_genv ge = let fundefs = ge.Genv.genv_fundefs in let methoddefs = ge.Genv.genv_methoddefs in let constructor = ge.Genv.genv_constructor in - bind (Obj.magic coq_Monad_optErr) (Obj.magic clinear_fundefs fundefs) + bind (Obj.magic coq_Monad_optErr) (Obj.magic cbasic_fundefs fundefs) (fun fundefs0 -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic clinear_methoddefs methoddefs) (fun methoddefs0 -> + (Obj.magic cbasic_methoddefs methoddefs) (fun methoddefs0 -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic clinear_constructor constructor) (fun constructor0 -> + (Obj.magic cbasic_constructor constructor) (fun constructor0 -> ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; Genv.genv_funcs = funcs; Genv.genv_methods = methods; Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; diff --git a/src/backend/extraction/Gen3.mli b/src/backend/extraction/Gen3.mli index 4648dec..b30c22d 100644 --- a/src/backend/extraction/Gen3.mli +++ b/src/backend/extraction/Gen3.mli @@ -1,49 +1,33 @@ -open AST open Ascii -open BinNums -open Coqlib open Datatypes open Globalenvs open Maps0 open Monad open OptErrMonad open Semantics -open Specif +open StmtCGraph open StmtClinear open String0 open Trees -val enumerate_rest : Semantics.code -> node list +val cbasic_stm : bool option -> node -> StmtCGraph.statement -> bblock -val labels_of_bblock : bblock -> label list +val cbasic_code : bool option -> StmtCGraph.code -> Semantics.code -val enumerate'_func : (node list, Semantics.code) sigT -> node list +val cbasic_function : + bool option -> StmtCGraph.coq_function -> Semantics.coq_function -val enumerate' : node list -> Semantics.code -> node list +val cbasic_fundef : + bool option -> StmtCGraph.coq_function -> Semantics.coq_function optErr -val enumerate : Semantics.coq_function -> node list +val cbasic_fundefs : + StmtCGraph.coq_function PTree.t -> Semantics.coq_function PTree.t optErr -val starts_with : label -> code -> bool +val cbasic_methoddefs : + StmtCGraph.coq_function option IntMap.t -> Semantics.coq_function option + IntMap.t optErr -val add_branch : label -> code -> code +val cbasic_constructor : + StmtCGraph.coq_function option -> Semantics.coq_function optErr -val clinear_block : bblock -> code -> code - -val clinear_node : Semantics.coq_function -> node -> code -> code - -val clinear_body : Semantics.coq_function -> node list -> code - -val clinear_function : Semantics.coq_function -> coq_function optErr - -val clinear_fundef : Semantics.coq_function -> coq_function optErr - -val clinear_fundefs : - Semantics.coq_function PTree.t -> coq_function PTree.t optErr - -val clinear_methoddefs : - Semantics.coq_function option IntMap.t -> coq_function option IntMap.t - optErr - -val clinear_constructor : Semantics.coq_function option -> coq_function optErr - -val clinear_genv : Semantics.genv -> genv optErr +val cbasic_genv : StmtCGraph.genv -> Semantics.genv optErr diff --git a/src/backend/extraction/Gen4.ml b/src/backend/extraction/Gen4.ml index 1cbe7d9..93b6b7b 100644 --- a/src/backend/extraction/Gen4.ml +++ b/src/backend/extraction/Gen4.ml @@ -1,275 +1,192 @@ open AST open Ascii +open BinNums +open Coqlib open Datatypes open Globalenvs -open Integers -open Labels -open List0 open Maps0 open Monad open OptErrMonad -open Options -open Semantics0 +open Semantics +open Specif open StmtClinear open String0 open Trees -(** val nodes_in_code : code -> label list **) +(** val enumerate_rest : Semantics.code -> node list **) -let rec nodes_in_code = function +let enumerate_rest c = + PTree.xkeys c Coq_xH + +(** val labels_of_bblock : bblock -> label list **) + +let rec labels_of_bblock = function | Coq_nil -> Coq_nil -| Coq_cons (s, rest) -> - let more_nodes = nodes_in_code rest in +| Coq_cons (s, b') -> + (match s with + | Sjump n -> Coq_cons (n, (labels_of_bblock b')) + | Sjumpi (_, n) -> Coq_cons (n, (labels_of_bblock b')) + | _ -> labels_of_bblock b') + +(** val enumerate'_func : (node list, Semantics.code) sigT -> node list **) + +let rec enumerate'_func x = + let todo = projT1 x in + let c = projT2 x in + let enumerate'0 = fun todo0 c0 -> enumerate'_func (Coq_existT (todo0, c0)) + in + (match todo with + | Coq_nil -> enumerate_rest c + | Coq_cons (n, ns) -> + let filtered_var = PTree.get n c in + (match filtered_var with + | Some b -> + Coq_cons (n, + (enumerate'0 (app (labels_of_bblock b) ns) (PTree.remove n c))) + | None -> enumerate'0 ns c)) + +(** val enumerate' : node list -> Semantics.code -> node list **) + +let enumerate' todo c = + enumerate'_func (Coq_existT (todo, c)) + +(** val enumerate : Semantics.coq_function -> node list **) + +let enumerate f = + enumerate' (Coq_cons (f.fn_entrypoint, Coq_nil)) f.Semantics.fn_code + +(** val starts_with : label -> code -> bool **) + +let starts_with lbl = function +| Coq_nil -> Coq_false +| Coq_cons (s, _) -> (match s with - | Slabel lbl -> Coq_cons (lbl, more_nodes) - | _ -> more_nodes) + | Slabel lbl' -> proj_sumbool (peq lbl lbl') + | _ -> Coq_false) -(** val allocate_labels : fd_label -> code -> label_map -> label_map **) +(** val add_branch : label -> code -> code **) -let allocate_labels fn c base = - let nodes = nodes_in_code c in - xallocate_labels (map (elt_key fn) nodes) - (xallocate_labels (map (call_elt_key fn) nodes) base) +let add_branch s k = + match starts_with s k with + | Coq_true -> k + | Coq_false -> Coq_cons ((Sjump s), k) -(** val allocate_labels_fundef : - fd_label -> coq_function -> label_map -> label_map **) +(** val clinear_block : bblock -> code -> code **) -let allocate_labels_fundef fn_l f base = - allocate_label (key (Coq_inl fn_l)) (allocate_labels fn_l f.fn_code base) +let rec clinear_block b k = + match b with + | Coq_nil -> k + | Coq_cons (i, b') -> + (match i with + | Slabel _ -> clinear_block b' k + | Sjump n -> add_branch n k + | _ -> Coq_cons (i, (clinear_block b' k))) -(** val xallocate_labels_fun : - (label, coq_function) prod list -> label_map -> label_map **) +(** val clinear_node : Semantics.coq_function -> node -> code -> code **) -let rec xallocate_labels_fun elts base = - match elts with - | Coq_nil -> base - | Coq_cons (p, rest) -> - let Coq_pair (fn, fd) = p in - let rest_allocated = xallocate_labels_fun rest base in - allocate_labels_fundef (Lfun fn) fd rest_allocated +let clinear_node f pc k = + match PTree.get pc f.Semantics.fn_code with + | Some b -> Coq_cons ((Slabel pc), (clinear_block b k)) + | None -> k -(** val allocate_labels_fun : - coq_function PTree.t -> label_map -> label_map **) +(** val clinear_body : Semantics.coq_function -> node list -> code **) -let allocate_labels_fun defs base = - xallocate_labels_fun (PTree.elements defs) base +let clinear_body f enum = + list_fold_right (clinear_node f) enum Coq_nil -(** val allocate_labels_methods : - Int.int list -> coq_function option IntMap.t -> label_map -> label_map **) +(** val clinear_function : Semantics.coq_function -> coq_function optErr **) -let rec allocate_labels_methods sigs defs base = - match sigs with - | Coq_nil -> base - | Coq_cons (sig0, rest) -> - let rest_allocated = allocate_labels_methods rest defs base in - (match IntMap.get sig0 defs with - | Some fd -> - allocate_label (key (Coq_inl (Lmultiplexer sig0))) - (allocate_labels_fundef (Lmethod sig0) fd rest_allocated) - | None -> rest_allocated) +let clinear_function f = + let enum = enumerate f in + ret (Obj.magic coq_Monad_optErr) { fn_return = f.Semantics.fn_return; + fn_params = f.Semantics.fn_params; fn_temps = f.Semantics.fn_temps; + fn_locals = f.Semantics.fn_locals; fn_code = + (add_branch f.fn_entrypoint (clinear_body f enum)) } -(** val allocate_labels_constructor : - coq_function option -> label_map -> label_map **) +(** val clinear_fundef : Semantics.coq_function -> coq_function optErr **) -let allocate_labels_constructor optc base = - match optc with - | Some fd -> allocate_labels_fundef Lconstructor fd base - | None -> base +let clinear_fundef = + clinear_function -(** val allocate_labels_ge : genv -> label_map **) +(** val clinear_fundefs : + Semantics.coq_function PTree.t -> coq_function PTree.t optErr **) -let allocate_labels_ge ge = - let sigs = ge.Genv.genv_methods in - let fundefs = ge.Genv.genv_fundefs in - let methoddefs = ge.Genv.genv_methoddefs in - let constructor = ge.Genv.genv_constructor in - allocate_label (key (Coq_inl Lbody)) - (allocate_labels_constructor constructor - (allocate_labels_methods sigs methoddefs - (allocate_labels_fun fundefs empty_label_map))) - -(** val clabeled_stm : - fd_label -> label_map -> statement -> statement optErr **) - -let clabeled_stm fn lm s = match s with -| Scall (rv, dst, args, retlbl) -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic label_key lm (Coq_inl (Lfun dst))) (fun dst' -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic label_key lm (call_node_within fn retlbl)) (fun retlbl' -> - ret (Obj.magic coq_Monad_optErr) (Scall (rv, dst', args, retlbl')))) -| Slabel lbl -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic label_key lm (node_within fn lbl)) (fun l -> - ret (Obj.magic coq_Monad_optErr) (Slabel l)) -| Sjump lbl -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic label_key lm (node_within fn lbl)) (fun lbl' -> - ret (Obj.magic coq_Monad_optErr) (Sjump lbl')) -| Sjumpi (e, lbl) -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic label_key lm (node_within fn lbl)) (fun lbl' -> - ret (Obj.magic coq_Monad_optErr) (Sjumpi (e, lbl'))) -| Stransfer (a, v, fail) -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic label_key lm (node_within fn fail)) (fun f -> - ret (Obj.magic coq_Monad_optErr) (Stransfer (a, v, f))) -| Scallmethod (a, r, s0, v, args, fail) -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic label_key lm (node_within fn fail)) (fun f -> - ret (Obj.magic coq_Monad_optErr) (Scallmethod (a, r, s0, v, args, f))) -| _ -> ret (Obj.magic coq_Monad_optErr) s - -(** val clabeled_code : fd_label -> label_map -> code -> code optErr **) - -let clabeled_code fn lm c = - map_error (clabeled_stm fn lm) c - -(** val fn_is_method : fd_label -> bool **) - -let fn_is_method = function -| Lfun _ -> Coq_false -| _ -> Coq_true - -(** val clabeled_function_body : label -> fd_label -> code -> code **) - -let clabeled_function_body l fn c = - Coq_cons ((Slabel l), (Coq_cons ((Sfetchargs (fn_is_method fn)), (Coq_cons - (Sintro, c))))) - -(** val clabeled_function : - label_map -> fd_label -> coq_function -> coq_function optErr **) - -let clabeled_function lm fn f = - bind (Obj.magic coq_Monad_optErr) (Obj.magic clabeled_code fn lm f.fn_code) - (fun c -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic label_key lm (Coq_inl fn)) - (fun l -> - let body = clabeled_function_body l fn c in - ret (Obj.magic coq_Monad_optErr) { fn_return = f.fn_return; fn_params = - f.fn_params; fn_temps = f.fn_temps; fn_locals = f.fn_locals; - fn_code = body })) - -(** val clabeled_fundef : - label_map -> fd_label -> coq_function -> coq_function optErr **) - -let clabeled_fundef = - clabeled_function - -(** val clabeled_fundef_f : - label_map -> label -> coq_function -> (label, coq_function) prod optErr **) - -let clabeled_fundef_f lm name fd = - bind (Obj.magic coq_Monad_optErr) - (Obj.magic clabeled_fundef lm (Lfun name) fd) (fun f -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic label_key lm (Coq_inl (Lfun name))) (fun l -> - ret (Obj.magic coq_Monad_optErr) (Coq_pair (l, f)))) +let clinear_fundefs t0 = + transl_tree clinear_fundef t0 -(** val clabeled_fundefs : - label_map -> coq_function PTree.t -> coq_function PTree.t optErr **) +(** val clinear_methoddefs : + Semantics.coq_function option IntMap.t -> coq_function option IntMap.t + optErr **) -let clabeled_fundefs lm t0 = - transl_tree_keys_move (clabeled_fundef_f lm) t0 +let clinear_methoddefs methods = + transl_map clinear_fundef methods -(** val xclabeled_methods : - label_map -> Int.int list -> coq_function option IntMap.t -> coq_function - option IntMap.t optErr **) +(** val clinear_constructor : + Semantics.coq_function option -> coq_function optErr **) -let rec xclabeled_methods lm sigs methods = - match sigs with - | Coq_nil -> ret (Obj.magic coq_Monad_optErr) (IntMap.init None) - | Coq_cons (sig0, rest) -> - bind (Obj.magic coq_Monad_optErr) (xclabeled_methods lm rest methods) - (fun crest -> - bind (Obj.magic coq_Monad_optErr) - (fromOption (IntMap.get sig0 (Obj.magic methods)) (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))) (fun meth -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic clabeled_fundef lm (Lmethod sig0) meth) (fun b -> - ret (Obj.magic coq_Monad_optErr) (IntMap.set sig0 (Some b) crest)))) - -(** val clabeled_methods : - label_map -> Int.int list -> coq_function option IntMap.t -> coq_function - option IntMap.t optErr **) - -let clabeled_methods = - xclabeled_methods - -(** val clabeled_constructor : - label_map -> coq_function option -> coq_function optErr **) - -let clabeled_constructor lm = function +let clinear_constructor = function | Some c -> - bind (Obj.magic coq_Monad_optErr) (clabeled_fundef lm Lconstructor c) - (fun f -> ret (Obj.magic coq_Monad_optErr) f) + bind (Obj.magic coq_Monad_optErr) (clinear_fundef c) (fun f -> + ret (Obj.magic coq_Monad_optErr) f) | None -> Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), EmptyString)))))))))))))))))))))))))))))))))))))))))) + Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -(** val clabeled_genv : label_map -> genv -> genv optErr **) +(** val clinear_genv : Semantics.genv -> genv optErr **) -let clabeled_genv lm ge = +let clinear_genv ge = let vars = ge.Genv.genv_vars in let funcs = ge.Genv.genv_funcs in let methods = ge.Genv.genv_methods in @@ -277,170 +194,14 @@ let clabeled_genv lm ge = let fundefs = ge.Genv.genv_fundefs in let methoddefs = ge.Genv.genv_methoddefs in let constructor = ge.Genv.genv_constructor in - (match label_functions fundefs with - | Coq_true -> - (match label_methods methoddefs with - | Coq_true -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic clabeled_fundefs lm fundefs) (fun fundefs0 -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic clabeled_methods lm methods methoddefs) - (fun methoddefs0 -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic clabeled_constructor lm constructor) - (fun constructor0 -> - ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; - Genv.genv_funcs = funcs; Genv.genv_methods = methods; - Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; - Genv.genv_methoddefs = methoddefs0; Genv.genv_constructor = - (Some constructor0) }))) - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - -(** val clabeled_program : genv -> program optErr **) - -let clabeled_program ge = - let lm = allocate_labels_ge ge in - bind (Obj.magic coq_Monad_optErr) (Obj.magic clabeled_genv lm ge) - (fun cge -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic clinear_fundefs fundefs) + (fun fundefs0 -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic label_key lm (Coq_inl Lbody)) (fun body -> - ret (Obj.magic coq_Monad_optErr) (Coq_pair (cge, body)))) + (Obj.magic clinear_methoddefs methoddefs) (fun methoddefs0 -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic clinear_constructor constructor) (fun constructor0 -> + ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; + Genv.genv_funcs = funcs; Genv.genv_methods = methods; + Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; + Genv.genv_methoddefs = methoddefs0; Genv.genv_constructor = (Some + constructor0) }))) diff --git a/src/backend/extraction/Gen4.mli b/src/backend/extraction/Gen4.mli index d16acbd..4648dec 100644 --- a/src/backend/extraction/Gen4.mli +++ b/src/backend/extraction/Gen4.mli @@ -1,70 +1,49 @@ open AST open Ascii +open BinNums +open Coqlib open Datatypes open Globalenvs -open Integers -open Labels -open List0 open Maps0 open Monad open OptErrMonad -open Options -open Semantics0 +open Semantics +open Specif open StmtClinear open String0 open Trees -val nodes_in_code : code -> label list +val enumerate_rest : Semantics.code -> node list -val allocate_labels : fd_label -> code -> label_map -> label_map +val labels_of_bblock : bblock -> label list -val allocate_labels_fundef : - fd_label -> coq_function -> label_map -> label_map +val enumerate'_func : (node list, Semantics.code) sigT -> node list -val xallocate_labels_fun : - (label, coq_function) prod list -> label_map -> label_map +val enumerate' : node list -> Semantics.code -> node list -val allocate_labels_fun : coq_function PTree.t -> label_map -> label_map +val enumerate : Semantics.coq_function -> node list -val allocate_labels_methods : - Int.int list -> coq_function option IntMap.t -> label_map -> label_map +val starts_with : label -> code -> bool -val allocate_labels_constructor : - coq_function option -> label_map -> label_map +val add_branch : label -> code -> code -val allocate_labels_ge : genv -> label_map +val clinear_block : bblock -> code -> code -val clabeled_stm : fd_label -> label_map -> statement -> statement optErr +val clinear_node : Semantics.coq_function -> node -> code -> code -val clabeled_code : fd_label -> label_map -> code -> code optErr +val clinear_body : Semantics.coq_function -> node list -> code -val fn_is_method : fd_label -> bool +val clinear_function : Semantics.coq_function -> coq_function optErr -val clabeled_function_body : label -> fd_label -> code -> code +val clinear_fundef : Semantics.coq_function -> coq_function optErr -val clabeled_function : - label_map -> fd_label -> coq_function -> coq_function optErr +val clinear_fundefs : + Semantics.coq_function PTree.t -> coq_function PTree.t optErr -val clabeled_fundef : - label_map -> fd_label -> coq_function -> coq_function optErr +val clinear_methoddefs : + Semantics.coq_function option IntMap.t -> coq_function option IntMap.t + optErr -val clabeled_fundef_f : - label_map -> label -> coq_function -> (label, coq_function) prod optErr +val clinear_constructor : Semantics.coq_function option -> coq_function optErr -val clabeled_fundefs : - label_map -> coq_function PTree.t -> coq_function PTree.t optErr - -val xclabeled_methods : - label_map -> Int.int list -> coq_function option IntMap.t -> coq_function - option IntMap.t optErr - -val clabeled_methods : - label_map -> Int.int list -> coq_function option IntMap.t -> coq_function - option IntMap.t optErr - -val clabeled_constructor : - label_map -> coq_function option -> coq_function optErr - -val clabeled_genv : label_map -> genv -> genv optErr - -val clabeled_program : genv -> program optErr +val clinear_genv : Semantics.genv -> genv optErr diff --git a/src/backend/extraction/Gen5.ml b/src/backend/extraction/Gen5.ml index b3de164..a5d582f 100644 --- a/src/backend/extraction/Gen5.ml +++ b/src/backend/extraction/Gen5.ml @@ -1,1156 +1,238 @@ +open AST open Ascii -open BinNums -open Cop -open Ctypes open Datatypes -open ExpCintptr -open ExpMiniC open Globalenvs +open Integers +open Labels +open List0 open Maps0 open Monad open OptErrMonad -open StackEnv -open StmtCintptr -open StmtClocal +open Options +open Semantics0 +open StmtClinear open String0 open Trees -(** val spE : ExpCintptr.expr **) +(** val nodes_in_code : code -> label list **) -let spE = - Emload ((ExpCintptr.Econst_int256 (sp, (Tpointer (Coq_mem, (Tpointer - (Coq_mem, Tvoid)))))), (Tpointer (Coq_mem, Tvoid))) +let rec nodes_in_code = function +| Coq_nil -> Coq_nil +| Coq_cons (s, rest) -> + let more_nodes = nodes_in_code rest in + (match s with + | Slabel lbl -> Coq_cons (lbl, more_nodes) + | _ -> more_nodes) -(** val offsetE : coq_function -> ExpCintptr.expr **) +(** val allocate_labels : fd_label -> code -> label_map -> label_map **) -let offsetE f = - ExpCintptr.Econst_int256 ((frame_size f.fn_locals), (Tint (I256, Unsigned))) +let allocate_labels fn c base = + let nodes = nodes_in_code c in + xallocate_labels (map (elt_key fn) nodes) + (xallocate_labels (map (call_elt_key fn) nodes) base) -(** val pushS : coq_function -> StmtCintptr.statement **) +(** val allocate_labels_fundef : + fd_label -> coq_function -> label_map -> label_map **) -let pushS f = - StmtCintptr.Smassign ((ExpCintptr.Econst_int256 (sp, (Tpointer (Coq_mem, - (Tpointer (Coq_mem, Tvoid)))))), (ExpCintptr.Ebinop (Oadd, spE, - (offsetE f), (Tpointer (Coq_mem, Tvoid))))) +let allocate_labels_fundef fn_l f base = + allocate_label (key (Coq_inl fn_l)) (allocate_labels fn_l f.fn_code base) -(** val popS : coq_function -> StmtCintptr.statement **) +(** val xallocate_labels_fun : + (label, coq_function) prod list -> label_map -> label_map **) -let popS f = - StmtCintptr.Smassign ((ExpCintptr.Econst_int256 (sp, (Tpointer (Coq_mem, - (Tpointer (Coq_mem, Tvoid)))))), (ExpCintptr.Ebinop (Osub, spE, - (offsetE f), (Tpointer (Coq_mem, Tvoid))))) +let rec xallocate_labels_fun elts base = + match elts with + | Coq_nil -> base + | Coq_cons (p, rest) -> + let Coq_pair (fn, fd) = p in + let rest_allocated = xallocate_labels_fun rest base in + allocate_labels_fundef (Lfun fn) fd rest_allocated -(** val cintptr_expr : coq_function -> expr -> ExpCintptr.expr optErr **) +(** val allocate_labels_fun : + coq_function PTree.t -> label_map -> label_map **) -let rec cintptr_expr f = function -| Econst_int (_, _) -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -| Econst_int256 (i, t0) -> - ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Econst_int256 (i, t0)) -| Evar (id, t0) -> - let locals = f.fn_locals in - let fl = mkfieldlist locals in - (match offset fl id with - | Some off -> - let foE = ExpCintptr.Econst_int256 ((frame_size locals), (Tint (I256, - Unsigned))) - in - let baseE = ExpCintptr.Ebinop (Osub, spE, foE, (Tpointer (Coq_mem, - (Tstruct (Coq_xH, fl))))) - in - let offE = ExpCintptr.Econst_int256 (off, (Tint (I256, Unsigned))) in - ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ebinop (Oadd, baseE, offE, - t0)) - | None -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), EmptyString))))))))))))))))))))))))))))))))))))))))))) -| Eglob (_, _) -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -| Etempvar (i, t0) -> - ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Etempvar (i, t0)) -| Ederef (e0, t0) -> - bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e0) (fun e' -> - match ExpCintptr.typeof e' with - | Tpointer (p, _) -> - (match p with - | Coq_mem -> ret (Obj.magic coq_Monad_optErr) (Emload (e', t0)) - | Coq_stor -> ret (Obj.magic coq_Monad_optErr) (Esload (e', t0)) - | Coq_call -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | _ -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -| Eaddr (e0, t0) -> - bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e0) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Eaddr (e', t0))) -| Eunop (o, e0, t0) -> - (match o with - | Osha_1 -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | _ -> - bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e0) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Eunop (o, e', t0)))) -| Ebinop (o, e1, e2, t0) -> - (match o with - | Osha_2 -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | _ -> - bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e1) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e2) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ebinop (o, e1', e2', - t0))))) -| Efield (e0, id, t0) -> - bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e0) (fun e' -> - match ExpCintptr.typeof e' with - | Tpointer (p, t1) -> - (match p with - | Coq_mem -> - (match t1 with - | Tstruct (_, fl) -> - (match offset fl id with - | Some off -> - let offE = ExpCintptr.Econst_int256 (off, (Tint (I256, - Unsigned))) - in - ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ebinop (Oadd, e', - offE, t0)) - | None -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | _ -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | _ -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | _ -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -| Eindex (e1, e2, t0) -> - bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e1) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e2) (fun e2' -> - match ExpCintptr.typeof e1' with - | Tpointer (p, t1) -> - (match p with - | Coq_mem -> - (match t1 with - | Tarray (t', _) -> - let sizeE = ExpCintptr.Econst_int256 ((sizeof t'), (Tint (I256, - Unsigned))) - in - let offE = ExpCintptr.Ebinop (Omul, e2', sizeE, (Tint (I256, - Unsigned))) - in - ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ebinop (Oadd, e1', - offE, t0)) - | _ -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | _ -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | _ -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, +let allocate_labels_fun defs base = + xallocate_labels_fun (PTree.elements defs) base + +(** val allocate_labels_methods : + Int.int list -> coq_function option IntMap.t -> label_map -> label_map **) + +let rec allocate_labels_methods sigs defs base = + match sigs with + | Coq_nil -> base + | Coq_cons (sig0, rest) -> + let rest_allocated = allocate_labels_methods rest defs base in + (match IntMap.get sig0 defs with + | Some fd -> + allocate_label (key (Coq_inl (Lmultiplexer sig0))) + (allocate_labels_fundef (Lmethod sig0) fd rest_allocated) + | None -> rest_allocated) + +(** val allocate_labels_constructor : + coq_function option -> label_map -> label_map **) + +let allocate_labels_constructor optc base = + match optc with + | Some fd -> allocate_labels_fundef Lconstructor fd base + | None -> base + +(** val allocate_labels_ge : genv -> label_map **) + +let allocate_labels_ge ge = + let sigs = ge.Genv.genv_methods in + let fundefs = ge.Genv.genv_fundefs in + let methoddefs = ge.Genv.genv_methoddefs in + let constructor = ge.Genv.genv_constructor in + allocate_label (key (Coq_inl Lbody)) + (allocate_labels_constructor constructor + (allocate_labels_methods sigs methoddefs + (allocate_labels_fun fundefs empty_label_map))) + +(** val clabeled_stm : + fd_label -> label_map -> statement -> statement optErr **) + +let clabeled_stm fn lm s = match s with +| Scall (rv, dst, args, retlbl) -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic label_key lm (Coq_inl (Lfun dst))) (fun dst' -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic label_key lm (call_node_within fn retlbl)) (fun retlbl' -> + ret (Obj.magic coq_Monad_optErr) (Scall (rv, dst', args, retlbl')))) +| Slabel lbl -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic label_key lm (node_within fn lbl)) (fun l -> + ret (Obj.magic coq_Monad_optErr) (Slabel l)) +| Sjump lbl -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic label_key lm (node_within fn lbl)) (fun lbl' -> + ret (Obj.magic coq_Monad_optErr) (Sjump lbl')) +| Sjumpi (e, lbl) -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic label_key lm (node_within fn lbl)) (fun lbl' -> + ret (Obj.magic coq_Monad_optErr) (Sjumpi (e, lbl'))) +| Stransfer (a, v, fail) -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic label_key lm (node_within fn fail)) (fun f -> + ret (Obj.magic coq_Monad_optErr) (Stransfer (a, v, f))) +| Scallmethod (a, r, s0, v, g, args, fail) -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic label_key lm (node_within fn fail)) (fun f -> + ret (Obj.magic coq_Monad_optErr) (Scallmethod (a, r, s0, v, g, args, f))) +| _ -> ret (Obj.magic coq_Monad_optErr) s + +(** val clabeled_code : fd_label -> label_map -> code -> code optErr **) + +let clabeled_code fn lm c = + map_error (clabeled_stm fn lm) c + +(** val fn_is_method : fd_label -> bool **) + +let fn_is_method = function +| Lfun _ -> Coq_false +| _ -> Coq_true + +(** val clabeled_function_body : label -> fd_label -> code -> code **) + +let clabeled_function_body l fn c = + Coq_cons ((Slabel l), (Coq_cons ((Sfetchargs (fn_is_method fn)), (Coq_cons + (Sintro, c))))) + +(** val clabeled_function : + label_map -> fd_label -> coq_function -> coq_function optErr **) + +let clabeled_function lm fn f = + bind (Obj.magic coq_Monad_optErr) (Obj.magic clabeled_code fn lm f.fn_code) + (fun c -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic label_key lm (Coq_inl fn)) + (fun l -> + let body = clabeled_function_body l fn c in + ret (Obj.magic coq_Monad_optErr) { fn_return = f.fn_return; fn_params = + f.fn_params; fn_temps = f.fn_temps; fn_locals = f.fn_locals; + fn_code = body })) + +(** val clabeled_fundef : + label_map -> fd_label -> coq_function -> coq_function optErr **) + +let clabeled_fundef = + clabeled_function + +(** val clabeled_fundef_f : + label_map -> label -> coq_function -> (label, coq_function) prod optErr **) + +let clabeled_fundef_f lm name fd = + bind (Obj.magic coq_Monad_optErr) + (Obj.magic clabeled_fundef lm (Lfun name) fd) (fun f -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic label_key lm (Coq_inl (Lfun name))) (fun l -> + ret (Obj.magic coq_Monad_optErr) (Coq_pair (l, f)))) + +(** val clabeled_fundefs : + label_map -> coq_function PTree.t -> coq_function PTree.t optErr **) + +let clabeled_fundefs lm t0 = + transl_tree_keys_move (clabeled_fundef_f lm) t0 + +(** val xclabeled_methods : + label_map -> Int.int list -> coq_function option IntMap.t -> coq_function + option IntMap.t optErr **) + +let rec xclabeled_methods lm sigs methods = + match sigs with + | Coq_nil -> ret (Obj.magic coq_Monad_optErr) (IntMap.init None) + | Coq_cons (sig0, rest) -> + bind (Obj.magic coq_Monad_optErr) (xclabeled_methods lm rest methods) + (fun crest -> + bind (Obj.magic coq_Monad_optErr) + (fromOption (IntMap.get sig0 (Obj.magic methods)) (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -| Ecall0 (b, t0) -> - ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ecall0 (b, t0)) -| Ecall1 (b, e0, t0) -> - bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e0) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ecall1 (b, e', t0))) - -(** val cintptr_exprs : - coq_function -> expr list -> ExpCintptr.expr list optErr **) - -let rec cintptr_exprs f = function -| Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil -| Coq_cons (e, es0) -> - bind (Obj.magic coq_Monad_optErr) (cintptr_exprs f es0) (fun es' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (e', es')))) - -(** val cintptr_expr_opt : - coq_function -> expr option -> ExpCintptr.expr option optErr **) - -let cintptr_expr_opt f = function -| Some e -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (Some e')) -| None -> ret (Obj.magic coq_Monad_optErr) None - -(** val cintptr_stmt : - coq_function -> statement -> StmtCintptr.statement optErr **) - -let rec cintptr_stmt f = function -| Sskip -> ret (Obj.magic coq_Monad_optErr) StmtCintptr.Sskip -| Smassign (e1, e2) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e1) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e2) - (fun e2' -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Smassign (e1', e2')))) -| Ssassign (e1, e2) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e1) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e2) - (fun e2' -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Ssassign (e1', e2')))) -| Sset (i, e) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e) (fun e' -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Sset (i, e'))) -| Scall (rv, dest, args) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_exprs f args) - (fun args' -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Scall (rv, dest, args'))) -| Ssequence (s1, s2) -> - bind (Obj.magic coq_Monad_optErr) (cintptr_stmt f s1) (fun s1' -> - bind (Obj.magic coq_Monad_optErr) (cintptr_stmt f s2) (fun s2' -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Ssequence (s1', s2')))) -| Sifthenelse (e, s1, s2) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e) (fun e' -> - bind (Obj.magic coq_Monad_optErr) (cintptr_stmt f s1) (fun s1' -> - bind (Obj.magic coq_Monad_optErr) (cintptr_stmt f s2) (fun s2' -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Sifthenelse (e', s1', - s2'))))) -| Sloop s0 -> - bind (Obj.magic coq_Monad_optErr) (cintptr_stmt f s0) (fun s' -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Sloop s')) -| Sbreak -> ret (Obj.magic coq_Monad_optErr) StmtCintptr.Sbreak -| Sreturn idopt -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Ssequence ((popS f), - (StmtCintptr.Sreturn idopt))) -| Shash (e1, e2, eopt) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e1) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e2) - (fun e2' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr_opt f eopt) - (fun eopt0 -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Shash (e1', e2', eopt0))))) -| Stransfer (a, v) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f a) (fun a' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f v) (fun v' -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Stransfer (a', v')))) -| Scallmethod (a, rvs, sg, v, args) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f a) (fun a' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f v) (fun v' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_exprs f args) - (fun args' -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Scallmethod (a', rvs, - sg, v', args'))))) -| Slog (topics, args) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_exprs f topics) - (fun topics' -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_exprs f args) - (fun args' -> - ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Slog (topics', args')))) -| Srevert -> ret (Obj.magic coq_Monad_optErr) StmtCintptr.Srevert - -(** val cintptr_function : coq_function -> StmtCintptr.coq_function optErr **) - -let cintptr_function f = - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_stmt f f.fn_body) - (fun s -> - ret (Obj.magic coq_Monad_optErr) { StmtCintptr.fn_return = f.fn_return; - StmtCintptr.fn_params = f.fn_params; StmtCintptr.fn_temps = f.fn_temps; - StmtCintptr.fn_locals = f.fn_locals; StmtCintptr.fn_body = - (StmtCintptr.Ssequence ((pushS f), s)) }) - -(** val cintptr_fundefs : - coq_function PTree.t -> StmtCintptr.coq_function PTree.t optErr **) - -let rec cintptr_fundefs t0 = - transl_tree cintptr_function t0 + (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))) (fun meth -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic clabeled_fundef lm (Lmethod sig0) meth) (fun b -> + ret (Obj.magic coq_Monad_optErr) (IntMap.set sig0 (Some b) crest)))) -(** val cintptr_methods : - coq_function option IntMap.t -> StmtCintptr.coq_function option IntMap.t - optErr **) +(** val clabeled_methods : + label_map -> Int.int list -> coq_function option IntMap.t -> coq_function + option IntMap.t optErr **) -let rec cintptr_methods methods = - transl_map cintptr_function methods +let clabeled_methods = + xclabeled_methods -(** val cintptr_constructor : - coq_function option -> StmtCintptr.coq_function optErr **) +(** val clabeled_constructor : + label_map -> coq_function option -> coq_function optErr **) -let cintptr_constructor = function +let clabeled_constructor lm = function | Some c -> - bind (Obj.magic coq_Monad_optErr) (cintptr_function c) (fun f -> - ret (Obj.magic coq_Monad_optErr) f) + bind (Obj.magic coq_Monad_optErr) (clabeled_fundef lm Lconstructor c) + (fun f -> ret (Obj.magic coq_Monad_optErr) f) | None -> Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, @@ -1185,9 +267,9 @@ let cintptr_constructor = function Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), EmptyString)))))))))))))))))))))))))))))))))))))))))) -(** val cintptr_genv : genv -> StmtCintptr.genv optErr **) +(** val clabeled_genv : label_map -> genv -> genv optErr **) -let cintptr_genv ge = +let clabeled_genv lm ge = let vars = ge.Genv.genv_vars in let funcs = ge.Genv.genv_funcs in let methods = ge.Genv.genv_methods in @@ -1195,14 +277,170 @@ let cintptr_genv ge = let fundefs = ge.Genv.genv_fundefs in let methoddefs = ge.Genv.genv_methoddefs in let constructor = ge.Genv.genv_constructor in - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_fundefs fundefs) - (fun fundefs0 -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_methods methoddefs) - (fun methoddefs0 -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic cintptr_constructor constructor) (fun constructor0 -> - ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; - Genv.genv_funcs = funcs; Genv.genv_methods = methods; - Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; - Genv.genv_methoddefs = methoddefs0; Genv.genv_constructor = (Some - constructor0) }))) + (match label_functions fundefs with + | Coq_true -> + (match label_methods methoddefs with + | Coq_true -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic clabeled_fundefs lm fundefs) (fun fundefs0 -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic clabeled_methods lm methods methoddefs) + (fun methoddefs0 -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic clabeled_constructor lm constructor) + (fun constructor0 -> + ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; + Genv.genv_funcs = funcs; Genv.genv_methods = methods; + Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; + Genv.genv_methoddefs = methoddefs0; Genv.genv_constructor = + (Some constructor0) }))) + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + +(** val clabeled_program : genv -> program optErr **) + +let clabeled_program ge = + let lm = allocate_labels_ge ge in + bind (Obj.magic coq_Monad_optErr) (Obj.magic clabeled_genv lm ge) + (fun cge -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic label_key lm (Coq_inl Lbody)) (fun body -> + ret (Obj.magic coq_Monad_optErr) (Coq_pair (cge, body)))) diff --git a/src/backend/extraction/Gen5.mli b/src/backend/extraction/Gen5.mli index e25ea88..d16acbd 100644 --- a/src/backend/extraction/Gen5.mli +++ b/src/backend/extraction/Gen5.mli @@ -1,47 +1,70 @@ +open AST open Ascii -open BinNums -open Cop -open Ctypes open Datatypes -open ExpCintptr -open ExpMiniC open Globalenvs +open Integers +open Labels +open List0 open Maps0 open Monad open OptErrMonad -open StackEnv -open StmtCintptr -open StmtClocal +open Options +open Semantics0 +open StmtClinear open String0 open Trees -val spE : ExpCintptr.expr +val nodes_in_code : code -> label list -val offsetE : coq_function -> ExpCintptr.expr +val allocate_labels : fd_label -> code -> label_map -> label_map -val pushS : coq_function -> StmtCintptr.statement +val allocate_labels_fundef : + fd_label -> coq_function -> label_map -> label_map -val popS : coq_function -> StmtCintptr.statement +val xallocate_labels_fun : + (label, coq_function) prod list -> label_map -> label_map -val cintptr_expr : coq_function -> expr -> ExpCintptr.expr optErr +val allocate_labels_fun : coq_function PTree.t -> label_map -> label_map -val cintptr_exprs : coq_function -> expr list -> ExpCintptr.expr list optErr +val allocate_labels_methods : + Int.int list -> coq_function option IntMap.t -> label_map -> label_map -val cintptr_expr_opt : - coq_function -> expr option -> ExpCintptr.expr option optErr +val allocate_labels_constructor : + coq_function option -> label_map -> label_map -val cintptr_stmt : coq_function -> statement -> StmtCintptr.statement optErr +val allocate_labels_ge : genv -> label_map -val cintptr_function : coq_function -> StmtCintptr.coq_function optErr +val clabeled_stm : fd_label -> label_map -> statement -> statement optErr -val cintptr_fundefs : - coq_function PTree.t -> StmtCintptr.coq_function PTree.t optErr +val clabeled_code : fd_label -> label_map -> code -> code optErr -val cintptr_methods : - coq_function option IntMap.t -> StmtCintptr.coq_function option IntMap.t - optErr +val fn_is_method : fd_label -> bool -val cintptr_constructor : - coq_function option -> StmtCintptr.coq_function optErr +val clabeled_function_body : label -> fd_label -> code -> code -val cintptr_genv : genv -> StmtCintptr.genv optErr +val clabeled_function : + label_map -> fd_label -> coq_function -> coq_function optErr + +val clabeled_fundef : + label_map -> fd_label -> coq_function -> coq_function optErr + +val clabeled_fundef_f : + label_map -> label -> coq_function -> (label, coq_function) prod optErr + +val clabeled_fundefs : + label_map -> coq_function PTree.t -> coq_function PTree.t optErr + +val xclabeled_methods : + label_map -> Int.int list -> coq_function option IntMap.t -> coq_function + option IntMap.t optErr + +val clabeled_methods : + label_map -> Int.int list -> coq_function option IntMap.t -> coq_function + option IntMap.t optErr + +val clabeled_constructor : + label_map -> coq_function option -> coq_function optErr + +val clabeled_genv : label_map -> genv -> genv optErr + +val clabeled_program : genv -> program optErr diff --git a/src/backend/extraction/Gen6.ml b/src/backend/extraction/Gen6.ml index 83028fd..d4cd5b2 100644 --- a/src/backend/extraction/Gen6.ml +++ b/src/backend/extraction/Gen6.ml @@ -1,1458 +1,1157 @@ -open AST open Ascii -open BinInt open BinNums open Cop open Ctypes open Datatypes open ExpCintptr -open ExpStacked -open FramesLabelsCintptr +open ExpMiniC open Globalenvs -open Integers -open List0 open Maps0 -open MemoryModel open Monad -open Nat0 open OptErrMonad -open Options -open Semantics1 -open StmtClinear -open StmtStacked +open StackEnv +open StmtCintptr +open StmtClocal open String0 -open TempModelLow open Trees -(** val stacked_expr : - nat PTree.t -> ExpCintptr.expr -> nat -> bool -> statement list optErr **) +(** val spE : ExpCintptr.expr **) -let rec stacked_expr temps e s_extra rvalue = - match e with - | Econst_int (_, _) -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | ExpCintptr.Econst_int256 (i, _) -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Srvalue (Econst_int256 i)), - Coq_nil)) - | ExpCintptr.Etempvar (i, _) -> - bind (Obj.magic coq_Monad_optErr) - (fromOption (PTree.get i (Obj.magic temps)) (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, +let spE = + Emload ((ExpCintptr.Econst_int256 (sp, (Tpointer (Coq_mem, (Tpointer + (Coq_mem, Tvoid)))))), (Tpointer (Coq_mem, Tvoid))) + +(** val offsetE : coq_function -> ExpCintptr.expr **) + +let offsetE f = + ExpCintptr.Econst_int256 ((frame_size f.fn_locals), (Tint (I256, Unsigned))) + +(** val pushS : coq_function -> StmtCintptr.statement **) + +let pushS f = + StmtCintptr.Smassign ((ExpCintptr.Econst_int256 (sp, (Tpointer (Coq_mem, + (Tpointer (Coq_mem, Tvoid)))))), (ExpCintptr.Ebinop (Oadd, spE, + (offsetE f), (Tpointer (Coq_mem, Tvoid))))) + +(** val popS : coq_function -> StmtCintptr.statement **) + +let popS f = + StmtCintptr.Smassign ((ExpCintptr.Econst_int256 (sp, (Tpointer (Coq_mem, + (Tpointer (Coq_mem, Tvoid)))))), (ExpCintptr.Ebinop (Osub, spE, + (offsetE f), (Tpointer (Coq_mem, Tvoid))))) + +(** val cintptr_expr : coq_function -> expr -> ExpCintptr.expr optErr **) + +let rec cintptr_expr f = function +| Econst_int (_, _) -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) +| Econst_int256 (i, t0) -> + ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Econst_int256 (i, t0)) +| Evar (id, t0) -> + let locals = f.fn_locals in + let fl = mkfieldlist locals in + (match offset fl id with + | Some off -> + let foE = ExpCintptr.Econst_int256 ((frame_size locals), (Tint (I256, + Unsigned))) + in + let baseE = ExpCintptr.Ebinop (Osub, spE, foE, (Tpointer (Coq_mem, + (Tstruct (Coq_xH, fl))))) + in + let offE = ExpCintptr.Econst_int256 (off, (Tint (I256, Unsigned))) in + ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ebinop (Oadd, baseE, offE, + t0)) + | None -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), EmptyString))))))))))))))))))))))))))))))))))))))))))) +| Eglob (_, _) -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) +| Etempvar (i, t0) -> + ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Etempvar (i, t0)) +| Ederef (e0, t0) -> + bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e0) (fun e' -> + match ExpCintptr.typeof e' with + | Tpointer (p, _) -> + (match p with + | Coq_mem -> ret (Obj.magic coq_Monad_optErr) (Emload (e', t0)) + | Coq_stor -> ret (Obj.magic coq_Monad_optErr) (Esload (e', t0)) + | Coq_call -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | _ -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) +| Eaddr (e0, t0) -> + bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e0) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Eaddr (e', t0))) +| Eunop (o, e0, t0) -> + (match o with + | Osha_1 -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | _ -> + bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e0) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Eunop (o, e', t0)))) +| Ebinop (o, e1, e2, t0) -> + (match o with + | Osha_2 -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | _ -> + bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e1) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e2) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ebinop (o, e1', e2', + t0))))) +| Efield (e0, id, t0) -> + bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e0) (fun e' -> + match ExpCintptr.typeof e' with + | Tpointer (p, t1) -> + (match p with + | Coq_mem -> + (match t1 with + | Tstruct (_, fl) -> + (match offset fl id with + | Some off -> + let offE = ExpCintptr.Econst_int256 (off, (Tint (I256, + Unsigned))) + in + ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ebinop (Oadd, e', + offE, t0)) + | None -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | _ -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | _ -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | _ -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - (fun ind -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Srvalue (Etempvar - (add s_extra ind))), Coq_nil))) - | ExpCintptr.Esload (e0, _) -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e0 s_extra Coq_true) (fun e' -> - ret (Obj.magic coq_Monad_optErr) - (app e' (Coq_cons ((Srvalue Esload), Coq_nil)))) - | ExpCintptr.Emload (e0, _) -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e0 s_extra Coq_true) (fun e' -> - ret (Obj.magic coq_Monad_optErr) - (app e' (Coq_cons ((Srvalue Emload), Coq_nil)))) - | Eaddr (e0, _) -> - (match rvalue with - | Coq_true -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e0 s_extra Coq_false) (fun e' -> - ret (Obj.magic coq_Monad_optErr) e') - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | ExpCintptr.Eunop (o, e0, _) -> - (match rvalue with - | Coq_true -> - (match o with - | Onotbool -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e0 s_extra Coq_true) (fun e' -> - ret (Obj.magic coq_Monad_optErr) - (app e' (Coq_cons ((Srvalue (Eunop o)), Coq_nil)))) - | Onotint -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e0 s_extra Coq_true) (fun e' -> - ret (Obj.magic coq_Monad_optErr) - (app e' (Coq_cons ((Srvalue (Eunop o)), Coq_nil)))) - | Oneg -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e0 s_extra Coq_true) (fun e' -> - ret (Obj.magic coq_Monad_optErr) - (app e' (Coq_cons ((Srvalue (Eunop o)), Coq_nil)))) - | Osha_1 -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Coq_false -> - (match o with - | Osha_1 -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | _ -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | ExpCintptr.Ebinop (o, e1, e2, _) -> - (match o with - | Oadd -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Osub -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Omul -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Odiv -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Omod -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Oexp -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Oand -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Oor -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Oxor -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Oshl -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Oshr -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Oeq -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | One -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Olt -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Ogt -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Ole -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Oge -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> - ret (Obj.magic coq_Monad_optErr) - (app e2' - (app e1' (Coq_cons - ((match rvalue with - | Coq_true -> Srvalue (Ebinop (o, Coq_false)) - | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) - | Osha_2 -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | ExpCintptr.Ecall0 (b, _) -> - (match rvalue with - | Coq_true -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Srvalue (Ecall0 b)), - Coq_nil)) - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | ExpCintptr.Ecall1 (b, e0, _) -> - (match rvalue with - | Coq_true -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e0 s_extra Coq_true) (fun e' -> - ret (Obj.magic coq_Monad_optErr) - (app e' (Coq_cons ((Srvalue (Ecall1 b)), Coq_nil)))) - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - -(** val stacked_exprs : - nat PTree.t -> ExpCintptr.expr list -> nat -> statement list optErr **) - -let rec stacked_exprs temps es s_extra = - match es with - | Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil - | Coq_cons (e, rest) -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps e (add (length rest) s_extra) Coq_true) (fun e' -> - bind (Obj.magic coq_Monad_optErr) (stacked_exprs temps rest s_extra) - (fun rest' -> ret (Obj.magic coq_Monad_optErr) (app rest' e'))) - -(** val ident_indices : - nat PTree.t -> ident list -> nat -> nat list optErr **) - -let rec ident_indices temps a offset = - match a with - | Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil - | Coq_cons (i, rest) -> - bind (Obj.magic coq_Monad_optErr) - (fromOption (PTree.get i (Obj.magic temps)) (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))) - (fun t0 -> - bind (Obj.magic coq_Monad_optErr) (ident_indices temps rest offset) - (fun resti -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons - ((add (add t0 offset) (length rest)), resti)))) - -(** val set_indices : nat list -> statement list **) - -let set_indices inds = - map (fun x -> Sset x) inds - -(** val optident : nat PTree.t -> ident option -> statement optErr **) - -let optident temps = function -| Some b -> - bind (Obj.magic coq_Monad_optErr) - (fromOption (PTree.get b (Obj.magic temps)) (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))) - (fun t0 -> ret (Obj.magic coq_Monad_optErr) (Sset t0)) -| None -> ret (Obj.magic coq_Monad_optErr) Spop - -(** val return_type : StmtClinear.coq_function -> coq_type **) - -let return_type = - fn_return - -(** val toreturn : StmtClinear.coq_function -> bool option -> ret_type **) - -let toreturn f = function -| Some b -> - (match b with - | Coq_true -> - (match return_type f with - | Tvoid -> Tvoid_method - | Tarray (_, _) -> Terror - | Tstruct (_, _) -> Terror - | _ -> Tsome_method) - | Coq_false -> - (match return_type f with - | Tvoid -> Tvoid_fun - | Tarray (_, _) -> Terror - | Tstruct (_, _) -> Terror - | _ -> Tsome_fun)) -| None -> - (match return_type f with - | Tvoid -> Tvoid_constructor - | Tarray (_, _) -> Terror - | Tstruct (_, _) -> Terror - | _ -> Tsome_constructor) - -(** val zero_stm : statement **) - -let zero_stm = - Srvalue (Econst_int256 Int256.zero) - -(** val xzero_stms : nat -> statement list **) - -let rec xzero_stms = function -| O -> Coq_nil -| S n -> app (xzero_stms n) (Coq_cons (zero_stm, Coq_nil)) - -(** val zero_stms : nat -> statement list **) - -let zero_stms c = match c with -| O -> Coq_cons (Sskip, Coq_nil) -| S _ -> xzero_stms c - -(** val z_stm : coq_Z -> statement **) - -let z_stm z = - Srvalue (Econst_int256 (Int256.repr z)) - -(** val stacked_code : - nat PTree.t -> StmtClinear.coq_function -> StmtClinear.code -> code optErr **) - -let rec stacked_code temps f = function -| Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil -| Coq_cons (s, rest) -> - bind (Obj.magic coq_Monad_optErr) (stacked_code temps f rest) (fun trest -> - match s with - | StmtClinear.Ssassign (lv, rv) -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps lv (S O) Coq_false) (fun lv' -> - bind (Obj.magic coq_Monad_optErr) (stacked_expr temps rv O Coq_true) - (fun rv' -> - ret (Obj.magic coq_Monad_optErr) - (app rv' (app lv' (Coq_cons (Ssassign, trest)))))) - | StmtClinear.Smassign (lv, rv) -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps lv (S O) Coq_false) (fun lv' -> - bind (Obj.magic coq_Monad_optErr) (stacked_expr temps rv O Coq_true) - (fun rv' -> - ret (Obj.magic coq_Monad_optErr) - (app rv' (app lv' (Coq_cons (Smassign, trest)))))) - | StmtClinear.Sset (i, rv) -> - bind (Obj.magic coq_Monad_optErr) - (fromOption (PTree.get i (Obj.magic temps)) (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) +| Eindex (e1, e2, t0) -> + bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e1) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e2) (fun e2' -> + match ExpCintptr.typeof e1' with + | Tpointer (p, t1) -> + (match p with + | Coq_mem -> + (match t1 with + | Tarray (t', _) -> + let sizeE = ExpCintptr.Econst_int256 ((sizeof t'), (Tint (I256, + Unsigned))) + in + let offE = ExpCintptr.Ebinop (Omul, e2', sizeE, (Tint (I256, + Unsigned))) + in + ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ebinop (Oadd, e1', + offE, t0)) + | _ -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | _ -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | _ -> + Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + (Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - (fun t0 -> - bind (Obj.magic coq_Monad_optErr) (stacked_expr temps rv O Coq_true) - (fun rv' -> - ret (Obj.magic coq_Monad_optErr) - (app rv' (Coq_cons ((Sset t0), trest))))) - | Scall (rv, dest, args, retlbl) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic optident temps rv) - (fun t0 -> - bind (Obj.magic coq_Monad_optErr) (stacked_exprs temps args (S O)) - (fun args' -> - ret (Obj.magic coq_Monad_optErr) - (app (Coq_cons ((Spushlabel (Lreturn retlbl)), Coq_nil)) - (app args' - (app (Coq_cons ((Spushlabel (Lcall dest)), (Coq_cons - (Sjump_call, (Coq_cons ((Slabel retlbl), Coq_nil)))))) - (app (Coq_cons (t0, Coq_nil)) trest)))))) - | Sreturn rv -> - (match rv with - | Some i -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps (ExpCintptr.Etempvar (i, (return_type f))) O - Coq_true) (fun rv' -> - ret (Obj.magic coq_Monad_optErr) (app rv' trest)) - | None -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (Spushvoid, trest))) - | StmtClinear.Sdone method0 -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Sdone - ((length (all_temps ftype (Obj.magic f))), (toreturn f method0))), - trest)) - | StmtClinear.Slabel lbl -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Slabel lbl), trest)) - | Sjump lbl -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Spushlabel (Linternal - lbl)), (Coq_cons (Sjump_internal, trest)))) - | StmtClinear.Sjumpi (cond, lbl) -> - bind (Obj.magic coq_Monad_optErr) (stacked_expr temps cond O Coq_true) - (fun cond' -> - ret (Obj.magic coq_Monad_optErr) - (app cond' (Coq_cons ((Spushlabel (Linternal lbl)), (Coq_cons - (Sjumpi, trest)))))) - | StmtClinear.Shash (ex1, ex2, exo) -> - bind (Obj.magic coq_Monad_optErr) (stacked_expr temps ex1 O Coq_true) - (fun ex1' -> - bind (Obj.magic coq_Monad_optErr) (stacked_expr temps ex2 O Coq_true) - (fun ex2' -> - bind (Obj.magic coq_Monad_optErr) - (match exo with - | Some ex -> stacked_expr temps ex O Coq_true - | None -> ret (Obj.magic coq_Monad_optErr) ex2') (fun exo' -> - bind (Obj.magic coq_Monad_optErr) - (match exo with - | Some _ -> - Success (Coq_cons ((Srvalue (Econst_int256 - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO Coq_xH)))))))))), Coq_nil)) - | None -> - Success (Coq_cons ((Srvalue (Econst_int256 - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO - Coq_xH))))))))), Coq_nil))) (fun len -> - ret (Obj.magic coq_Monad_optErr) - (app ex2' - (app (Coq_cons ((Srvalue Emload), (Coq_cons ((Srvalue - (Econst_int256 - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO (Coq_xO (Coq_xO (Coq_xO Coq_xH)))))))))))), - (Coq_cons (Smassign, Coq_nil)))))) - (app - (match exo with - | Some _ -> - app exo' (Coq_cons ((Srvalue Emload), (Coq_cons - ((Srvalue (Econst_int256 - (Int256.repr - (Z.add (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO (Coq_xO (Coq_xO (Coq_xO - Coq_xH))))))))) (Zpos (Coq_xO (Coq_xO (Coq_xO - (Coq_xO (Coq_xO Coq_xH)))))))))), (Coq_cons - (Smassign, Coq_nil)))))) - | None -> Coq_nil) - (app len - (app (Coq_cons ((Srvalue (Econst_int256 - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO (Coq_xO (Coq_xO (Coq_xO Coq_xH)))))))))))), - Coq_nil)) - (app (Coq_cons (Shash, Coq_nil)) - (app ex1' - (app (Coq_cons (Smassign, Coq_nil)) trest)))))))))))) - | StmtClinear.Stransfer (a, v, fail) -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps a (S (S (S (S (S O))))) Coq_true) (fun a' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps v (S (S (S (S O)))) Coq_true) (fun v' -> - ret (Obj.magic coq_Monad_optErr) - (app (Coq_cons (zero_stm, (Coq_cons (zero_stm, (Coq_cons - (zero_stm, (Coq_cons (zero_stm, Coq_nil)))))))) - (app v' - (app a' - (app (Coq_cons (zero_stm, (Coq_cons (Stransfer, (Coq_cons - ((Srvalue (Eunop Onotbool)), (Coq_cons ((Spushlabel - (Linternal fail)), (Coq_cons (Sjumpi, Coq_nil)))))))))) - trest)))))) - | StmtClinear.Scallmethod (a, rvs, sg, v, args, fail) -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps a (S (S (S (S (S O))))) Coq_true) (fun a' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_expr temps v (S (S (S (S O)))) Coq_true) (fun v' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_exprs temps args (S (S (S (S (S (S (S O)))))))) - (fun args' -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic ident_indices temps rvs (S O)) (fun rv_inds -> - let retcount = length rvs in - let argcount = length args in - ret (Obj.magic coq_Monad_optErr) - (app (Coq_cons ((z_stm (arglen argcount)), (Coq_cons - ((z_stm (argpos retcount)), (Coq_cons - ((z_stm (retlen retcount)), (Coq_cons ((z_stm retpos), - Coq_nil)))))))) - (app v' - (app a' (Coq_cons (zero_stm, - (app args' (Coq_cons ((Scallmethod (sg, argcount, - retcount)), - (app (set_indices rv_inds) - (app (Coq_cons ((Srvalue (Eunop Onotbool)), - (Coq_cons ((Spushlabel (Linternal fail)), - (Coq_cons (Sjumpi, Coq_nil)))))) trest)))))))))))))) - | StmtClinear.Slog (topics, args) -> - bind (Obj.magic coq_Monad_optErr) (stacked_exprs temps topics O) - (fun topics' -> - bind (Obj.magic coq_Monad_optErr) - (stacked_exprs temps args (length topics)) (fun args' -> - ret (Obj.magic coq_Monad_optErr) - (app topics' - (app args' (Coq_cons ((Slog ((length topics), (length args))), - trest)))))) - | StmtClinear.Srevert -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (Srevert, trest)) - | StmtClinear.Sfetchargs fetch -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons - ((match fetch with - | Coq_true -> Sfetchargs (length (some_args ftype (Obj.magic f))) - | Coq_false -> Sskip), trest)) - | Sintro -> - let push_temps = zero_stms (length (some_temps ftype (Obj.magic f))) in - ret (Obj.magic coq_Monad_optErr) (app push_temps trest)) + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) +| Ecall0 (b, t0) -> + ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ecall0 (b, t0)) +| Ecall1 (b, e0, t0) -> + bind (Obj.magic coq_Monad_optErr) (cintptr_expr f e0) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (ExpCintptr.Ecall1 (b, e', t0))) -(** val allocate_temps : (ident, coq_type) prod list -> nat -> nat PTree.t **) +(** val cintptr_exprs : + coq_function -> expr list -> ExpCintptr.expr list optErr **) -let rec allocate_temps ids base = - match ids with - | Coq_nil -> PTree.empty - | Coq_cons (p, rest) -> - let Coq_pair (id, _) = p in - PTree.set id base (allocate_temps rest (S base)) +let rec cintptr_exprs f = function +| Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil +| Coq_cons (e, es0) -> + bind (Obj.magic coq_Monad_optErr) (cintptr_exprs f es0) (fun es' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (e', es')))) -(** val allocate_all_temps : (ident, coq_type) prod list -> nat PTree.t **) +(** val cintptr_expr_opt : + coq_function -> expr option -> ExpCintptr.expr option optErr **) -let allocate_all_temps ids = - allocate_temps ids O +let cintptr_expr_opt f = function +| Some e -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (Some e')) +| None -> ret (Obj.magic coq_Monad_optErr) None -(** val allocate_fn_temps : StmtClinear.coq_function -> nat PTree.t **) +(** val cintptr_stmt : + coq_function -> statement -> StmtCintptr.statement optErr **) -let allocate_fn_temps f = - allocate_all_temps (all_temps ftype (Obj.magic f)) +let rec cintptr_stmt f = function +| Sskip -> ret (Obj.magic coq_Monad_optErr) StmtCintptr.Sskip +| Smassign (e1, e2) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e1) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e2) + (fun e2' -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Smassign (e1', e2')))) +| Ssassign (e1, e2) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e1) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e2) + (fun e2' -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Ssassign (e1', e2')))) +| Sset (i, e) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e) (fun e' -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Sset (i, e'))) +| Scall (rv, dest, args) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_exprs f args) + (fun args' -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Scall (rv, dest, args'))) +| Ssequence (s1, s2) -> + bind (Obj.magic coq_Monad_optErr) (cintptr_stmt f s1) (fun s1' -> + bind (Obj.magic coq_Monad_optErr) (cintptr_stmt f s2) (fun s2' -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Ssequence (s1', s2')))) +| Sifthenelse (e, s1, s2) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e) (fun e' -> + bind (Obj.magic coq_Monad_optErr) (cintptr_stmt f s1) (fun s1' -> + bind (Obj.magic coq_Monad_optErr) (cintptr_stmt f s2) (fun s2' -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Sifthenelse (e', s1', + s2'))))) +| Sloop s0 -> + bind (Obj.magic coq_Monad_optErr) (cintptr_stmt f s0) (fun s' -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Sloop s')) +| Sbreak -> ret (Obj.magic coq_Monad_optErr) StmtCintptr.Sbreak +| Sreturn idopt -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Ssequence ((popS f), + (StmtCintptr.Sreturn idopt))) +| Shash (e1, e2, eopt) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e1) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f e2) + (fun e2' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr_opt f eopt) + (fun eopt0 -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Shash (e1', e2', eopt0))))) +| Stransfer (a, v) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f a) (fun a' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f v) (fun v' -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Stransfer (a', v')))) +| Scallmethod (a, rvs, sg, v, g, args) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f a) (fun a' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr f v) (fun v' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_expr_opt f g) + (fun g' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_exprs f args) + (fun args' -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Scallmethod (a', rvs, + sg, v', g', args')))))) +| Slog (topics, args) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_exprs f topics) + (fun topics' -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_exprs f args) + (fun args' -> + ret (Obj.magic coq_Monad_optErr) (StmtCintptr.Slog (topics', args')))) +| Srevert -> ret (Obj.magic coq_Monad_optErr) StmtCintptr.Srevert -(** val stacked_function : - StmtClinear.coq_function -> StmtStacked.coq_function optErr **) +(** val cintptr_function : coq_function -> StmtCintptr.coq_function optErr **) -let stacked_function f = - let temps = allocate_fn_temps f in - bind (Obj.magic coq_Monad_optErr) - (Obj.magic stacked_code temps f f.StmtClinear.fn_code) (fun c -> - ret (Obj.magic coq_Monad_optErr) c) +let cintptr_function f = + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_stmt f f.fn_body) + (fun s -> + ret (Obj.magic coq_Monad_optErr) { StmtCintptr.fn_return = f.fn_return; + StmtCintptr.fn_params = f.fn_params; StmtCintptr.fn_temps = f.fn_temps; + StmtCintptr.fn_locals = f.fn_locals; StmtCintptr.fn_body = + (StmtCintptr.Ssequence ((pushS f), s)) }) -(** val stacked_fundef : - StmtClinear.coq_function -> StmtStacked.coq_function optErr **) +(** val cintptr_fundefs : + coq_function PTree.t -> StmtCintptr.coq_function PTree.t optErr **) -let stacked_fundef = - stacked_function +let rec cintptr_fundefs t0 = + transl_tree cintptr_function t0 -(** val stacked_fundefs : - StmtClinear.coq_function PTree.t -> StmtStacked.coq_function PTree.t +(** val cintptr_methods : + coq_function option IntMap.t -> StmtCintptr.coq_function option IntMap.t optErr **) -let stacked_fundefs t0 = - transl_tree stacked_fundef t0 - -(** val stacked_methods : - StmtClinear.coq_function option IntMap.t -> StmtStacked.coq_function - option IntMap.t optErr **) - -let stacked_methods methods = - transl_map stacked_fundef methods +let rec cintptr_methods methods = + transl_map cintptr_function methods -(** val stacked_constructor : - StmtClinear.coq_function option -> StmtStacked.coq_function optErr **) +(** val cintptr_constructor : + coq_function option -> StmtCintptr.coq_function optErr **) -let stacked_constructor = function +let cintptr_constructor = function | Some c -> - bind (Obj.magic coq_Monad_optErr) (stacked_fundef c) (fun f -> + bind (Obj.magic coq_Monad_optErr) (cintptr_function c) (fun f -> ret (Obj.magic coq_Monad_optErr) f) | None -> Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, @@ -1488,9 +1187,9 @@ let stacked_constructor = function Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), EmptyString)))))))))))))))))))))))))))))))))))))))))) -(** val stacked_genv : StmtClinear.genv -> genv optErr **) +(** val cintptr_genv : genv -> StmtCintptr.genv optErr **) -let stacked_genv ge = +let cintptr_genv ge = let vars = ge.Genv.genv_vars in let funcs = ge.Genv.genv_funcs in let methods = ge.Genv.genv_methods in @@ -1498,146 +1197,14 @@ let stacked_genv ge = let fundefs = ge.Genv.genv_fundefs in let methoddefs = ge.Genv.genv_methoddefs in let constructor = ge.Genv.genv_constructor in - bind (Obj.magic coq_Monad_optErr) (Obj.magic stacked_fundefs fundefs) + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_fundefs fundefs) (fun fundefs0 -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic stacked_methods methoddefs) + bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_methods methoddefs) (fun methoddefs0 -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic stacked_constructor constructor) (fun constructor0 -> + (Obj.magic cintptr_constructor constructor) (fun constructor0 -> ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; Genv.genv_funcs = funcs; Genv.genv_methods = methods; Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; Genv.genv_methoddefs = methoddefs0; Genv.genv_constructor = (Some constructor0) }))) - -(** val stacked_program : StmtClinear.program -> program optErr **) - -let stacked_program = function -| Coq_pair (ge, body) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic stacked_genv ge) (fun cge -> - match label_functions cge.Genv.genv_fundefs with - | Coq_true -> - (match label_methods cge.Genv.genv_methoddefs with - | Coq_true -> ret (Obj.magic coq_Monad_optErr) (Coq_pair (cge, body)) - | Coq_false -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Coq_false -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) diff --git a/src/backend/extraction/Gen6.mli b/src/backend/extraction/Gen6.mli index 204a37f..e25ea88 100644 --- a/src/backend/extraction/Gen6.mli +++ b/src/backend/extraction/Gen6.mli @@ -1,78 +1,47 @@ -open AST open Ascii -open BinInt open BinNums open Cop open Ctypes open Datatypes open ExpCintptr -open ExpStacked -open FramesLabelsCintptr +open ExpMiniC open Globalenvs -open Integers -open List0 open Maps0 -open MemoryModel open Monad -open Nat0 open OptErrMonad -open Options -open Semantics1 -open StmtClinear -open StmtStacked +open StackEnv +open StmtCintptr +open StmtClocal open String0 -open TempModelLow open Trees -val stacked_expr : - nat PTree.t -> ExpCintptr.expr -> nat -> bool -> statement list optErr +val spE : ExpCintptr.expr -val stacked_exprs : - nat PTree.t -> ExpCintptr.expr list -> nat -> statement list optErr +val offsetE : coq_function -> ExpCintptr.expr -val ident_indices : nat PTree.t -> ident list -> nat -> nat list optErr +val pushS : coq_function -> StmtCintptr.statement -val set_indices : nat list -> statement list +val popS : coq_function -> StmtCintptr.statement -val optident : nat PTree.t -> ident option -> statement optErr +val cintptr_expr : coq_function -> expr -> ExpCintptr.expr optErr -val return_type : StmtClinear.coq_function -> coq_type +val cintptr_exprs : coq_function -> expr list -> ExpCintptr.expr list optErr -val toreturn : StmtClinear.coq_function -> bool option -> ret_type +val cintptr_expr_opt : + coq_function -> expr option -> ExpCintptr.expr option optErr -val zero_stm : statement +val cintptr_stmt : coq_function -> statement -> StmtCintptr.statement optErr -val xzero_stms : nat -> statement list +val cintptr_function : coq_function -> StmtCintptr.coq_function optErr -val zero_stms : nat -> statement list +val cintptr_fundefs : + coq_function PTree.t -> StmtCintptr.coq_function PTree.t optErr -val z_stm : coq_Z -> statement +val cintptr_methods : + coq_function option IntMap.t -> StmtCintptr.coq_function option IntMap.t + optErr -val stacked_code : - nat PTree.t -> StmtClinear.coq_function -> StmtClinear.code -> code optErr +val cintptr_constructor : + coq_function option -> StmtCintptr.coq_function optErr -val allocate_temps : (ident, coq_type) prod list -> nat -> nat PTree.t - -val allocate_all_temps : (ident, coq_type) prod list -> nat PTree.t - -val allocate_fn_temps : StmtClinear.coq_function -> nat PTree.t - -val stacked_function : - StmtClinear.coq_function -> StmtStacked.coq_function optErr - -val stacked_fundef : - StmtClinear.coq_function -> StmtStacked.coq_function optErr - -val stacked_fundefs : - StmtClinear.coq_function PTree.t -> StmtStacked.coq_function PTree.t optErr - -val stacked_methods : - StmtClinear.coq_function option IntMap.t -> StmtStacked.coq_function option - IntMap.t optErr - -val stacked_constructor : - StmtClinear.coq_function option -> StmtStacked.coq_function optErr - -val stacked_genv : StmtClinear.genv -> genv optErr - -val stacked_program : StmtClinear.program -> program optErr +val cintptr_genv : genv -> StmtCintptr.genv optErr diff --git a/src/backend/extraction/Gen7.ml b/src/backend/extraction/Gen7.ml index 30c6d25..627aa7e 100644 --- a/src/backend/extraction/Gen7.ml +++ b/src/backend/extraction/Gen7.ml @@ -1,333 +1,1500 @@ open AST open Ascii +open BinInt open BinNums +open Cop +open Ctypes open Datatypes +open ExpCintptr open ExpStacked +open FramesLabelsCintptr open Globalenvs open Integers -open LowValues +open List0 open Maps0 open MemoryModel open Monad +open Nat0 open OptErrMonad -open Semantics2 -open StmtExpressionless +open Options +open Semantics1 +open StmtClinear open StmtStacked open String0 +open TempModelLow open Trees -(** val expressionless_expr : expr -> StmtExpressionless.statement **) +(** val stacked_expr : + nat PTree.t -> ExpCintptr.expr -> nat -> bool -> statement list optErr **) -let expressionless_expr = function -| Econst_int256 i -> Spush (Coq_inl (Vint i)) -| Eglob i -> Spush (Coq_inl (Vhash (Vint (Int256.repr (Zpos i))))) -| Etempvar n -> Sdup n -| Emload -> Smload -| Esload -> Ssload -| Eunop o -> Sunop o -| Ebinop (o, s) -> Sbinop (o, s) -| Ecall0 b -> Scall0 b -| Ecall1 b -> Scall1 b - -(** val pops : nat -> StmtExpressionless.statement list **) - -let rec pops = function -| O -> Coq_nil -| S m -> Coq_cons (StmtExpressionless.Spop, (pops m)) - -(** val cleanup : nat -> StmtExpressionless.statement list **) - -let cleanup n = match n with -| O -> Coq_nil -| S m -> Coq_cons ((Sswap m), (pops n)) - -(** val expressionless_rt : ExpStacked.ret_type -> ret_type optErr **) - -let expressionless_rt = function -| Tvoid_fun -> ret (Obj.magic coq_Monad_optErr) Tfun -| ExpStacked.Tvoid_method -> ret (Obj.magic coq_Monad_optErr) Tvoid_method -| Terror -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -| Tsome_fun -> ret (Obj.magic coq_Monad_optErr) Tfun -| ExpStacked.Tsome_method -> ret (Obj.magic coq_Monad_optErr) Tsome_method -| _ -> ret (Obj.magic coq_Monad_optErr) Tconstructor - -(** val fetch_args : - function_kind -> nat -> nat -> StmtExpressionless.statement list **) - -let rec fetch_args fk count base = - match count with - | O -> Coq_nil - | S n -> - app (fetch_args fk n (S base)) - (match fk with - | Coq_normalFunction -> - Coq_cons ((Spush (Coq_inl (Vint (call_data_arg_location base)))), - (Coq_cons (Scalldataload, Coq_nil))) - | Coq_constructorFunction -> - Coq_cons ((Sconstructordataload base), Coq_nil)) - -(** val extract_lbl : typed_label -> label **) - -let extract_lbl = function -| Linternal l' -> l' -| Lcall l' -> l' -| Lreturn l' -> l' - -(** val expressionless_stm : - statement -> function_kind -> StmtExpressionless.statement list optErr **) - -let expressionless_stm s kind = - match s with - | Sskip -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Sskip, - Coq_nil)) - | Srvalue e -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((expressionless_expr e), - Coq_nil)) - | Slvalue e -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((expressionless_expr e), - Coq_nil)) - | Spushvoid -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Spush (Coq_inl Vunit)), +let rec stacked_expr temps e s_extra rvalue = + match e with + | Econst_int (_, _) -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | ExpCintptr.Econst_int256 (i, _) -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Srvalue (Econst_int256 i)), Coq_nil)) - | Spop -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Spop, - Coq_nil)) - | Ssassign -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (Ssstore, Coq_nil)) - | Smassign -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (Smstore, Coq_nil)) - | Sset n -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Sswap n), (Coq_cons - (StmtExpressionless.Spop, Coq_nil)))) - | Sdone (n, rt) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic expressionless_rt rt) - (fun rt' -> + | ExpCintptr.Etempvar (i, _) -> + bind (Obj.magic coq_Monad_optErr) + (fromOption (PTree.get i (Obj.magic temps)) (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + (fun ind -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Srvalue (Etempvar + (add s_extra ind))), Coq_nil))) + | ExpCintptr.Esload (e0, _) -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e0 s_extra Coq_true) (fun e' -> ret (Obj.magic coq_Monad_optErr) - (app (cleanup n) (Coq_cons ((StmtExpressionless.Sdone rt'), Coq_nil)))) - | Spushlabel l -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Spush (Coq_inr - (extract_lbl l))), Coq_nil)) - | Slabel l -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((StmtExpressionless.Slabel - l), Coq_nil)) - | Sjump_call -> - (match kind with - | Coq_normalFunction -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (Sjump, Coq_nil)) - | Coq_constructorFunction -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + (app e' (Coq_cons ((Srvalue Esload), Coq_nil)))) + | ExpCintptr.Emload (e0, _) -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e0 s_extra Coq_true) (fun e' -> + ret (Obj.magic coq_Monad_optErr) + (app e' (Coq_cons ((Srvalue Emload), Coq_nil)))) + | Eaddr (e0, _) -> + (match rvalue with + | Coq_true -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e0 s_extra Coq_false) (fun e' -> + ret (Obj.magic coq_Monad_optErr) e') + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | ExpCintptr.Eunop (o, e0, _) -> + (match rvalue with + | Coq_true -> + (match o with + | Onotbool -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e0 s_extra Coq_true) (fun e' -> + ret (Obj.magic coq_Monad_optErr) + (app e' (Coq_cons ((Srvalue (Eunop o)), Coq_nil)))) + | Onotint -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e0 s_extra Coq_true) (fun e' -> + ret (Obj.magic coq_Monad_optErr) + (app e' (Coq_cons ((Srvalue (Eunop o)), Coq_nil)))) + | Oneg -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e0 s_extra Coq_true) (fun e' -> + ret (Obj.magic coq_Monad_optErr) + (app e' (Coq_cons ((Srvalue (Eunop o)), Coq_nil)))) + | Osha_1 -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Coq_false -> + (match o with + | Osha_1 -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | _ -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | ExpCintptr.Ebinop (o, e1, e2, _) -> + (match o with + | Oadd -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Osub -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Omul -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Odiv -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Omod -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Oexp -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Oand -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Oor -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Oxor -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Oshl -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Oshr -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Oeq -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | One -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Olt -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Ogt -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Ole -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Oge -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e1 (S s_extra) rvalue) (fun e1' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e2 s_extra Coq_true) (fun e2' -> + ret (Obj.magic coq_Monad_optErr) + (app e2' + (app e1' (Coq_cons + ((match rvalue with + | Coq_true -> Srvalue (Ebinop (o, Coq_false)) + | Coq_false -> Slvalue (Ebinop (o, Coq_false))), Coq_nil)))))) + | Osha_2 -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | ExpCintptr.Ecall0 (b, _) -> + (match rvalue with + | Coq_true -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Srvalue (Ecall0 b)), + Coq_nil)) + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | ExpCintptr.Ecall1 (b, e0, _) -> + (match rvalue with + | Coq_true -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e0 s_extra Coq_true) (fun e' -> + ret (Obj.magic coq_Monad_optErr) + (app e' (Coq_cons ((Srvalue (Ecall1 b)), Coq_nil)))) + | Coq_false -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Sjump_internal -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (Sjump, Coq_nil)) - | Sjumpi -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Sjumpi, - Coq_nil)) - | Shash -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Shash, - Coq_nil)) - | Stransfer -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Stransfer, - Coq_nil)) - | Scallmethod (i, a, r) -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons - ((StmtExpressionless.Scallmethod (i, a, r)), Coq_nil)) - | Slog (n, m) -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((StmtExpressionless.Slog (n, - m)), Coq_nil)) - | Srevert -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Srevert, - Coq_nil)) - | Sfetchargs n -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Sskip, - (fetch_args kind n O))) + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -(** val expressionless_code : - code -> function_kind -> StmtExpressionless.code optErr **) +(** val stacked_exprs : + nat PTree.t -> ExpCintptr.expr list -> nat -> statement list optErr **) -let rec expressionless_code c kind = - match c with +let rec stacked_exprs temps es s_extra = + match es with | Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil - | Coq_cons (s, rest) -> - bind (Obj.magic coq_Monad_optErr) (expressionless_stm s kind) (fun s' -> - bind (Obj.magic coq_Monad_optErr) (expressionless_code rest kind) - (fun rest' -> ret (Obj.magic coq_Monad_optErr) (app s' rest'))) + | Coq_cons (e, rest) -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps e (add (length rest) s_extra) Coq_true) (fun e' -> + bind (Obj.magic coq_Monad_optErr) (stacked_exprs temps rest s_extra) + (fun rest' -> ret (Obj.magic coq_Monad_optErr) (app rest' e'))) -(** val expressionless_function : - function_kind -> coq_function -> StmtExpressionless.coq_function optErr **) +(** val ident_indices : + nat PTree.t -> ident list -> nat -> nat list optErr **) -let expressionless_function kind f = +let rec ident_indices temps a offset = + match a with + | Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil + | Coq_cons (i, rest) -> + bind (Obj.magic coq_Monad_optErr) + (fromOption (PTree.get i (Obj.magic temps)) (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))) + (fun t0 -> + bind (Obj.magic coq_Monad_optErr) (ident_indices temps rest offset) + (fun resti -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons + ((add (add t0 offset) (length rest)), resti)))) + +(** val set_indices : nat list -> statement list **) + +let set_indices inds = + map (fun x -> Sset x) inds + +(** val load_returns : nat -> statement list **) + +let rec load_returns = function +| O -> Coq_nil +| S n -> + app (Coq_cons ((Srvalue (Econst_int256 + (Int256.repr (Z.add retpos (retlen n))))), (Coq_cons ((Srvalue Emload), + Coq_nil)))) (load_returns n) + +(** val optident : nat PTree.t -> ident option -> statement optErr **) + +let optident temps = function +| Some b -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic expressionless_code (fn_code f) kind) (fun c -> + (fromOption (PTree.get b (Obj.magic temps)) (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))) + (fun t0 -> ret (Obj.magic coq_Monad_optErr) (Sset t0)) +| None -> ret (Obj.magic coq_Monad_optErr) Spop + +(** val return_type : StmtClinear.coq_function -> coq_type **) + +let return_type = + fn_return + +(** val toreturn : StmtClinear.coq_function -> bool option -> ret_type **) + +let toreturn f = function +| Some b -> + (match b with + | Coq_true -> + (match return_type f with + | Tvoid -> Tvoid_method + | Tarray (_, _) -> Terror + | Tstruct (_, _) -> Terror + | _ -> Tsome_method) + | Coq_false -> + (match return_type f with + | Tvoid -> Tvoid_fun + | Tarray (_, _) -> Terror + | Tstruct (_, _) -> Terror + | _ -> Tsome_fun)) +| None -> + (match return_type f with + | Tvoid -> Tvoid_constructor + | Tarray (_, _) -> Terror + | Tstruct (_, _) -> Terror + | _ -> Tsome_constructor) + +(** val zero_stm : statement **) + +let zero_stm = + Srvalue (Econst_int256 Int256.zero) + +(** val xzero_stms : nat -> statement list **) + +let rec xzero_stms = function +| O -> Coq_nil +| S n -> app (xzero_stms n) (Coq_cons (zero_stm, Coq_nil)) + +(** val zero_stms : nat -> statement list **) + +let zero_stms c = match c with +| O -> Coq_cons (Sskip, Coq_nil) +| S _ -> xzero_stms c + +(** val z_stm : coq_Z -> statement **) + +let z_stm z = + Srvalue (Econst_int256 (Int256.repr z)) + +(** val stacked_code : + nat PTree.t -> StmtClinear.coq_function -> StmtClinear.code -> code optErr **) + +let rec stacked_code temps f = function +| Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil +| Coq_cons (s, rest) -> + bind (Obj.magic coq_Monad_optErr) (stacked_code temps f rest) (fun trest -> + match s with + | StmtClinear.Ssassign (lv, rv) -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps lv (S O) Coq_false) (fun lv' -> + bind (Obj.magic coq_Monad_optErr) (stacked_expr temps rv O Coq_true) + (fun rv' -> + ret (Obj.magic coq_Monad_optErr) + (app rv' (app lv' (Coq_cons (Ssassign, trest)))))) + | StmtClinear.Smassign (lv, rv) -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps lv (S O) Coq_false) (fun lv' -> + bind (Obj.magic coq_Monad_optErr) (stacked_expr temps rv O Coq_true) + (fun rv' -> + ret (Obj.magic coq_Monad_optErr) + (app rv' (app lv' (Coq_cons (Smassign, trest)))))) + | StmtClinear.Sset (i, rv) -> + bind (Obj.magic coq_Monad_optErr) + (fromOption (PTree.get i (Obj.magic temps)) (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + (fun t0 -> + bind (Obj.magic coq_Monad_optErr) (stacked_expr temps rv O Coq_true) + (fun rv' -> + ret (Obj.magic coq_Monad_optErr) + (app rv' (Coq_cons ((Sset t0), trest))))) + | Scall (rv, dest, args, retlbl) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic optident temps rv) + (fun t0 -> + bind (Obj.magic coq_Monad_optErr) (stacked_exprs temps args (S O)) + (fun args' -> + ret (Obj.magic coq_Monad_optErr) + (app (Coq_cons ((Spushlabel (Lreturn retlbl)), Coq_nil)) + (app args' + (app (Coq_cons ((Spushlabel (Lcall dest)), (Coq_cons + (Sjump_call, (Coq_cons ((Slabel retlbl), Coq_nil)))))) + (app (Coq_cons (t0, Coq_nil)) trest)))))) + | Sreturn rv -> + (match rv with + | Some i -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps (ExpCintptr.Etempvar (i, (return_type f))) O + Coq_true) (fun rv' -> + ret (Obj.magic coq_Monad_optErr) (app rv' trest)) + | None -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (Spushvoid, trest))) + | StmtClinear.Sdone method0 -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Sdone + ((length (all_temps ftype (Obj.magic f))), (toreturn f method0))), + trest)) + | StmtClinear.Slabel lbl -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Slabel lbl), trest)) + | Sjump lbl -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Spushlabel (Linternal + lbl)), (Coq_cons (Sjump_internal, trest)))) + | StmtClinear.Sjumpi (cond, lbl) -> + bind (Obj.magic coq_Monad_optErr) (stacked_expr temps cond O Coq_true) + (fun cond' -> + ret (Obj.magic coq_Monad_optErr) + (app cond' (Coq_cons ((Spushlabel (Linternal lbl)), (Coq_cons + (Sjumpi, trest)))))) + | StmtClinear.Shash (ex1, ex2, exo) -> + bind (Obj.magic coq_Monad_optErr) (stacked_expr temps ex1 O Coq_true) + (fun ex1' -> + bind (Obj.magic coq_Monad_optErr) (stacked_expr temps ex2 O Coq_true) + (fun ex2' -> + bind (Obj.magic coq_Monad_optErr) + (match exo with + | Some ex -> stacked_expr temps ex O Coq_true + | None -> ret (Obj.magic coq_Monad_optErr) ex2') (fun exo' -> + bind (Obj.magic coq_Monad_optErr) + (match exo with + | Some _ -> + Success (Coq_cons ((Srvalue (Econst_int256 + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO Coq_xH)))))))))), Coq_nil)) + | None -> + Success (Coq_cons ((Srvalue (Econst_int256 + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO + Coq_xH))))))))), Coq_nil))) (fun len -> + ret (Obj.magic coq_Monad_optErr) + (app ex2' + (app (Coq_cons ((Srvalue Emload), (Coq_cons ((Srvalue + (Econst_int256 + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO (Coq_xO (Coq_xO (Coq_xO Coq_xH)))))))))))), + (Coq_cons (Smassign, Coq_nil)))))) + (app + (match exo with + | Some _ -> + app exo' (Coq_cons ((Srvalue Emload), (Coq_cons + ((Srvalue (Econst_int256 + (Int256.repr + (Z.add (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO (Coq_xO (Coq_xO (Coq_xO + Coq_xH))))))))) (Zpos (Coq_xO (Coq_xO (Coq_xO + (Coq_xO (Coq_xO Coq_xH)))))))))), (Coq_cons + (Smassign, Coq_nil)))))) + | None -> Coq_nil) + (app len + (app (Coq_cons ((Srvalue (Econst_int256 + (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO + (Coq_xO (Coq_xO (Coq_xO (Coq_xO Coq_xH)))))))))))), + Coq_nil)) + (app (Coq_cons (Shash, Coq_nil)) + (app ex1' + (app (Coq_cons (Smassign, Coq_nil)) trest)))))))))))) + | StmtClinear.Stransfer (a, v, fail) -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps a (S (S (S (S (S O))))) Coq_true) (fun a' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps v (S (S (S (S O)))) Coq_true) (fun v' -> + ret (Obj.magic coq_Monad_optErr) + (app (Coq_cons (zero_stm, (Coq_cons (zero_stm, (Coq_cons + (zero_stm, (Coq_cons (zero_stm, Coq_nil)))))))) + (app v' + (app a' + (app (Coq_cons (zero_stm, (Coq_cons (Stransfer, (Coq_cons + ((Srvalue (Eunop Onotbool)), (Coq_cons ((Spushlabel + (Linternal fail)), (Coq_cons (Sjumpi, Coq_nil)))))))))) + trest)))))) + | StmtClinear.Scallmethod (a, rvs, sg, v, g, args, fail) -> + (match g with + | Some gas -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps gas (S (S (S (S (S (S O)))))) Coq_true) + (fun g' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps a (S (S (S (S (S O))))) Coq_true) (fun a' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps v (S (S (S (S O)))) Coq_true) (fun v' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_exprs temps (rev args) O) (fun args' -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic ident_indices temps rvs O) (fun rv_inds -> + let retcount = length rvs in + let argcount = length args in + ret (Obj.magic coq_Monad_optErr) + (app args' (Coq_cons ((Scallargs (sg, argcount, + retcount)), + (app (Coq_cons ((z_stm (retlen retcount)), (Coq_cons + ((z_stm retpos), (Coq_cons + ((z_stm (arglen argcount)), (Coq_cons + ((z_stm (argpos retcount)), Coq_nil)))))))) + (app v' + (app a' + (app g' + (app (Coq_cons ((Scallmethod Coq_true), + Coq_nil)) + (app (Coq_cons ((Srvalue (Eunop Onotbool)), + (Coq_cons ((Spushlabel (Linternal fail)), + (Coq_cons (Sjumpi, Coq_nil)))))) + (app (load_returns retcount) + (app (set_indices rv_inds) trest)))))))))))))))) + | None -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps a (S (S (S (S (S O))))) Coq_true) (fun a' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_expr temps v (S (S (S (S O)))) Coq_true) (fun v' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_exprs temps (rev args) O) (fun args' -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic ident_indices temps rvs O) (fun rv_inds -> + let retcount = length rvs in + let argcount = length args in + ret (Obj.magic coq_Monad_optErr) + (app args' (Coq_cons ((Scallargs (sg, argcount, + retcount)), + (app (Coq_cons ((z_stm (retlen retcount)), (Coq_cons + ((z_stm retpos), (Coq_cons ((z_stm (arglen argcount)), + (Coq_cons ((z_stm (argpos retcount)), Coq_nil)))))))) + (app v' + (app a' + (app (Coq_cons ((Scallmethod Coq_false), Coq_nil)) + (app (Coq_cons ((Srvalue (Eunop Onotbool)), + (Coq_cons ((Spushlabel (Linternal fail)), + (Coq_cons (Sjumpi, Coq_nil)))))) + (app (load_returns retcount) + (app (set_indices rv_inds) trest))))))))))))))) + | StmtClinear.Slog (topics, args) -> + bind (Obj.magic coq_Monad_optErr) (stacked_exprs temps topics O) + (fun topics' -> + bind (Obj.magic coq_Monad_optErr) + (stacked_exprs temps args (length topics)) (fun args' -> + ret (Obj.magic coq_Monad_optErr) + (app topics' + (app args' (Coq_cons ((Slog ((length topics), (length args))), + trest)))))) + | StmtClinear.Srevert -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (Srevert, trest)) + | StmtClinear.Sfetchargs fetch -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons + ((match fetch with + | Coq_true -> Sfetchargs (length (some_args ftype (Obj.magic f))) + | Coq_false -> Sskip), trest)) + | Sintro -> + let push_temps = zero_stms (length (some_temps ftype (Obj.magic f))) in + ret (Obj.magic coq_Monad_optErr) (app push_temps trest)) + +(** val allocate_temps : (ident, coq_type) prod list -> nat -> nat PTree.t **) + +let rec allocate_temps ids base = + match ids with + | Coq_nil -> PTree.empty + | Coq_cons (p, rest) -> + let Coq_pair (id, _) = p in + PTree.set id base (allocate_temps rest (S base)) + +(** val allocate_all_temps : (ident, coq_type) prod list -> nat PTree.t **) + +let allocate_all_temps ids = + allocate_temps ids O + +(** val allocate_fn_temps : StmtClinear.coq_function -> nat PTree.t **) + +let allocate_fn_temps f = + allocate_all_temps (all_temps ftype (Obj.magic f)) + +(** val stacked_function : + StmtClinear.coq_function -> StmtStacked.coq_function optErr **) + +let stacked_function f = + let temps = allocate_fn_temps f in + bind (Obj.magic coq_Monad_optErr) + (Obj.magic stacked_code temps f f.StmtClinear.fn_code) (fun c -> ret (Obj.magic coq_Monad_optErr) c) -(** val expressionless_fundefs : - coq_function PTree.t -> StmtExpressionless.coq_function PTree.t optErr **) +(** val stacked_fundef : + StmtClinear.coq_function -> StmtStacked.coq_function optErr **) -let expressionless_fundefs t0 = - transl_tree (expressionless_function Coq_normalFunction) t0 +let stacked_fundef = + stacked_function -(** val expressionless_methods : - coq_function option IntMap.t -> StmtExpressionless.coq_function option - IntMap.t optErr **) +(** val stacked_fundefs : + StmtClinear.coq_function PTree.t -> StmtStacked.coq_function PTree.t + optErr **) -let expressionless_methods methods = - transl_map (expressionless_function Coq_normalFunction) methods +let stacked_fundefs t0 = + transl_tree stacked_fundef t0 -(** val expressionless_constructor : - coq_function option -> StmtExpressionless.coq_function optErr **) +(** val stacked_methods : + StmtClinear.coq_function option IntMap.t -> StmtStacked.coq_function + option IntMap.t optErr **) -let expressionless_constructor = function +let stacked_methods methods = + transl_map stacked_fundef methods + +(** val stacked_constructor : + StmtClinear.coq_function option -> StmtStacked.coq_function optErr **) + +let stacked_constructor = function | Some c -> - bind (Obj.magic coq_Monad_optErr) - (expressionless_function Coq_constructorFunction c) (fun f -> + bind (Obj.magic coq_Monad_optErr) (stacked_fundef c) (fun f -> ret (Obj.magic coq_Monad_optErr) f) | None -> Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, @@ -363,9 +1530,9 @@ let expressionless_constructor = function Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), EmptyString)))))))))))))))))))))))))))))))))))))))))) -(** val expressionless_genv : genv -> StmtExpressionless.genv optErr **) +(** val stacked_genv : StmtClinear.genv -> genv optErr **) -let expressionless_genv ge = +let stacked_genv ge = let vars = ge.Genv.genv_vars in let funcs = ge.Genv.genv_funcs in let methods = ge.Genv.genv_methods in @@ -373,100 +1540,146 @@ let expressionless_genv ge = let fundefs = ge.Genv.genv_fundefs in let methoddefs = ge.Genv.genv_methoddefs in let constructor = ge.Genv.genv_constructor in - bind (Obj.magic coq_Monad_optErr) - (Obj.magic expressionless_fundefs fundefs) (fun fundefs0 -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic expressionless_methods methoddefs) (fun methoddefs0 -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic stacked_fundefs fundefs) + (fun fundefs0 -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic stacked_methods methoddefs) + (fun methoddefs0 -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic expressionless_constructor constructor) - (fun constructor0 -> + (Obj.magic stacked_constructor constructor) (fun constructor0 -> ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; Genv.genv_funcs = funcs; Genv.genv_methods = methods; Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; Genv.genv_methoddefs = methoddefs0; Genv.genv_constructor = (Some constructor0) }))) -(** val expressionless_program : - program -> StmtExpressionless.program optErr **) +(** val stacked_program : StmtClinear.program -> program optErr **) -let expressionless_program = function +let stacked_program = function | Coq_pair (ge, body) -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic expressionless_genv ge) - (fun cge -> - match label_verify cge with - | Coq_true -> ret (Obj.magic coq_Monad_optErr) (Coq_pair (cge, body)) + bind (Obj.magic coq_Monad_optErr) (Obj.magic stacked_genv ge) (fun cge -> + match label_functions cge.Genv.genv_fundefs with + | Coq_true -> + (match label_methods cge.Genv.genv_methoddefs with + | Coq_true -> ret (Obj.magic coq_Monad_optErr) (Coq_pair (cge, body)) + | Coq_false -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) | Coq_false -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) diff --git a/src/backend/extraction/Gen7.mli b/src/backend/extraction/Gen7.mli index cfbf553..c704823 100644 --- a/src/backend/extraction/Gen7.mli +++ b/src/backend/extraction/Gen7.mli @@ -1,53 +1,80 @@ open AST open Ascii +open BinInt open BinNums +open Cop +open Ctypes open Datatypes +open ExpCintptr open ExpStacked +open FramesLabelsCintptr open Globalenvs open Integers -open LowValues +open List0 open Maps0 open MemoryModel open Monad +open Nat0 open OptErrMonad -open Semantics2 -open StmtExpressionless +open Options +open Semantics1 +open StmtClinear open StmtStacked open String0 +open TempModelLow open Trees -val expressionless_expr : expr -> StmtExpressionless.statement +val stacked_expr : + nat PTree.t -> ExpCintptr.expr -> nat -> bool -> statement list optErr -val pops : nat -> StmtExpressionless.statement list +val stacked_exprs : + nat PTree.t -> ExpCintptr.expr list -> nat -> statement list optErr -val cleanup : nat -> StmtExpressionless.statement list +val ident_indices : nat PTree.t -> ident list -> nat -> nat list optErr -val expressionless_rt : ExpStacked.ret_type -> ret_type optErr +val set_indices : nat list -> statement list -val fetch_args : - function_kind -> nat -> nat -> StmtExpressionless.statement list +val load_returns : nat -> statement list -val extract_lbl : typed_label -> label +val optident : nat PTree.t -> ident option -> statement optErr -val expressionless_stm : - statement -> function_kind -> StmtExpressionless.statement list optErr +val return_type : StmtClinear.coq_function -> coq_type -val expressionless_code : - code -> function_kind -> StmtExpressionless.code optErr +val toreturn : StmtClinear.coq_function -> bool option -> ret_type -val expressionless_function : - function_kind -> coq_function -> StmtExpressionless.coq_function optErr +val zero_stm : statement -val expressionless_fundefs : - coq_function PTree.t -> StmtExpressionless.coq_function PTree.t optErr +val xzero_stms : nat -> statement list -val expressionless_methods : - coq_function option IntMap.t -> StmtExpressionless.coq_function option +val zero_stms : nat -> statement list + +val z_stm : coq_Z -> statement + +val stacked_code : + nat PTree.t -> StmtClinear.coq_function -> StmtClinear.code -> code optErr + +val allocate_temps : (ident, coq_type) prod list -> nat -> nat PTree.t + +val allocate_all_temps : (ident, coq_type) prod list -> nat PTree.t + +val allocate_fn_temps : StmtClinear.coq_function -> nat PTree.t + +val stacked_function : + StmtClinear.coq_function -> StmtStacked.coq_function optErr + +val stacked_fundef : + StmtClinear.coq_function -> StmtStacked.coq_function optErr + +val stacked_fundefs : + StmtClinear.coq_function PTree.t -> StmtStacked.coq_function PTree.t optErr + +val stacked_methods : + StmtClinear.coq_function option IntMap.t -> StmtStacked.coq_function option IntMap.t optErr -val expressionless_constructor : - coq_function option -> StmtExpressionless.coq_function optErr +val stacked_constructor : + StmtClinear.coq_function option -> StmtStacked.coq_function optErr -val expressionless_genv : genv -> StmtExpressionless.genv optErr +val stacked_genv : StmtClinear.genv -> genv optErr -val expressionless_program : program -> StmtExpressionless.program optErr +val stacked_program : StmtClinear.program -> program optErr diff --git a/src/backend/extraction/Gen8.ml b/src/backend/extraction/Gen8.ml index 68aceb9..985d1e0 100644 --- a/src/backend/extraction/Gen8.ml +++ b/src/backend/extraction/Gen8.ml @@ -1,385 +1,475 @@ open AST open Ascii -open Cop +open BinNums open Datatypes +open ExpStacked open Globalenvs open Integers -open Language -open List0 open LowValues open Maps0 +open MemoryModel open Monad open OptErrMonad -open Options -open Semantics3 -open StackEnv +open Semantics2 open StmtExpressionless +open StmtStacked open String0 -open Zpower +open Trees -(** val methodical_fundefs : coq_function PTree.t -> code **) +(** val expressionless_expr : expr -> StmtExpressionless.statement **) -let methodical_fundefs t0 = - flat_map fn_code (map snd (PTree.elements t0)) +let expressionless_expr = function +| Econst_int256 i -> Spush (Coq_inl (Vint i)) +| Eglob i -> Spush (Coq_inl (Vhash (Vint (Int256.repr (Zpos i))))) +| Etempvar n -> Sdup n +| Emload -> Smload +| Esload -> Ssload +| Eunop o -> Sunop o +| Ebinop (o, s) -> Sbinop (o, s) +| Ecall0 b -> Scall0 b +| Ecall1 b -> Scall1 b -(** val methodical_opt_function : coq_function option -> code **) +(** val pops : nat -> StmtExpressionless.statement list **) -let methodical_opt_function = function -| Some f -> fn_code f -| None -> Coq_nil +let rec pops = function +| O -> Coq_nil +| S m -> Coq_cons (StmtExpressionless.Spop, (pops m)) -(** val methodical_methods : coq_function option IntMap.t -> code **) +(** val cleanup : nat -> StmtExpressionless.statement list **) -let methodical_methods methods = - flat_map methodical_opt_function (map snd (PTree.elements (snd methods))) +let cleanup n = match n with +| O -> Coq_nil +| S m -> Coq_cons ((Sswap m), (pops n)) -(** val label_method_starts_with : coq_function -> label optErr **) +(** val expressionless_rt : ExpStacked.ret_type -> ret_type optErr **) -let label_method_starts_with m = - match fn_code m with - | Coq_nil -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Coq_cons (s, _) -> - (match s with - | Slabel l -> ret (Obj.magic coq_Monad_optErr) l - | _ -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, +let expressionless_rt = function +| Tvoid_fun -> ret (Obj.magic coq_Monad_optErr) Tfun +| ExpStacked.Tvoid_method -> ret (Obj.magic coq_Monad_optErr) Tvoid_method +| Terror -> + Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) +| Tsome_fun -> ret (Obj.magic coq_Monad_optErr) Tfun +| ExpStacked.Tsome_method -> ret (Obj.magic coq_Monad_optErr) Tsome_method +| _ -> ret (Obj.magic coq_Monad_optErr) Tconstructor + +(** val fetch_args : + function_kind -> nat -> nat -> StmtExpressionless.statement list **) + +let rec fetch_args fk count base = + match count with + | O -> Coq_nil + | S n -> + app (fetch_args fk n (S base)) + (match fk with + | Coq_normalFunction -> + Coq_cons ((Spush (Coq_inl (Vint (call_data_arg_location base)))), + (Coq_cons (Scalldataload, Coq_nil))) + | Coq_constructorFunction -> + Coq_cons ((Sconstructordataload base), Coq_nil)) + +(** val extract_lbl : typed_label -> label **) + +let extract_lbl = function +| Linternal l' -> l' +| Lcall l' -> l' +| Lreturn l' -> l' + +(** val expressionless_stm : + statement -> function_kind -> StmtExpressionless.statement list optErr **) + +let expressionless_stm s kind = + match s with + | Sskip -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Sskip, + Coq_nil)) + | Srvalue e -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((expressionless_expr e), + Coq_nil)) + | Slvalue e -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((expressionless_expr e), + Coq_nil)) + | Spushvoid -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Spush (Coq_inl Vunit)), + Coq_nil)) + | Spop -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Spop, + Coq_nil)) + | Ssassign -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (Ssstore, Coq_nil)) + | Smassign -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (Smstore, Coq_nil)) + | Sset n -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Sswap n), (Coq_cons + (StmtExpressionless.Spop, Coq_nil)))) + | Sdone (n, rt) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic expressionless_rt rt) + (fun rt' -> + ret (Obj.magic coq_Monad_optErr) + (app (cleanup n) (Coq_cons ((StmtExpressionless.Sdone rt'), Coq_nil)))) + | Spushlabel l -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Spush (Coq_inr + (extract_lbl l))), Coq_nil)) + | Slabel l -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((StmtExpressionless.Slabel + l), Coq_nil)) + | Sjump_call -> + (match kind with + | Coq_normalFunction -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (Sjump, Coq_nil)) + | Coq_constructorFunction -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Sjump_internal -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (Sjump, Coq_nil)) + | Sjumpi -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Sjumpi, + Coq_nil)) + | Shash -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Shash, + Coq_nil)) + | Stransfer -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Stransfer, + Coq_nil)) + | Scallargs (i, a, r) -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((StmtExpressionless.Scallargs + (i, a, r)), Coq_nil)) + | Scallmethod b -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons + ((StmtExpressionless.Scallmethod b), Coq_nil)) + | Slog (n, m) -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((StmtExpressionless.Slog (n, + m)), Coq_nil)) + | Srevert -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Srevert, + Coq_nil)) + | Sfetchargs n -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons (StmtExpressionless.Sskip, + (fetch_args kind n O))) -(** val sg_val : Int.int -> coq_val **) +(** val expressionless_code : + code -> function_kind -> StmtExpressionless.code optErr **) -let sg_val sg = - Vint (Int256.repr (Int.unsigned sg)) +let rec expressionless_code c kind = + match c with + | Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil + | Coq_cons (s, rest) -> + bind (Obj.magic coq_Monad_optErr) (expressionless_stm s kind) (fun s' -> + bind (Obj.magic coq_Monad_optErr) (expressionless_code rest kind) + (fun rest' -> ret (Obj.magic coq_Monad_optErr) (app s' rest'))) -(** val methodical_multiplexer_body : - Int.int list -> coq_function option IntMap.t -> code optErr **) +(** val expressionless_function : + function_kind -> coq_function -> StmtExpressionless.coq_function optErr **) -let rec methodical_multiplexer_body methods methoddefs = - match methods with - | Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil - | Coq_cons (sg, rest) -> - bind (Obj.magic coq_Monad_optErr) - (methodical_multiplexer_body rest methoddefs) (fun rest' -> - bind (Obj.magic coq_Monad_optErr) - (fromOption (IntMap.get sg (Obj.magic methoddefs)) (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - (fun m -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic label_method_starts_with m) (fun l -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Sdup O), (Coq_cons - ((Spush (Coq_inl (sg_val sg))), (Coq_cons ((Sbinop (Oeq, - Coq_false)), (Coq_cons ((Spush (Coq_inr l)), (Coq_cons (Sjumpi, - rest'))))))))))))) +let expressionless_function kind f = + bind (Obj.magic coq_Monad_optErr) + (Obj.magic expressionless_code (fn_code f) kind) (fun c -> + ret (Obj.magic coq_Monad_optErr) c) + +(** val expressionless_fundefs : + coq_function PTree.t -> StmtExpressionless.coq_function PTree.t optErr **) + +let expressionless_fundefs t0 = + transl_tree (expressionless_function Coq_normalFunction) t0 + +(** val expressionless_methods : + coq_function option IntMap.t -> StmtExpressionless.coq_function option + IntMap.t optErr **) + +let expressionless_methods methods = + transl_map (expressionless_function Coq_normalFunction) methods -(** val methodical_main : program -> code optErr **) +(** val expressionless_constructor : + coq_function option -> StmtExpressionless.coq_function optErr **) -let methodical_main p = - let ge = fst p in - let body = snd p in +let expressionless_constructor = function +| Some c -> + bind (Obj.magic coq_Monad_optErr) + (expressionless_function Coq_constructorFunction c) (fun f -> + ret (Obj.magic coq_Monad_optErr) f) +| None -> + Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), EmptyString)))))))))))))))))))))))))))))))))))))))))) + +(** val expressionless_genv : genv -> StmtExpressionless.genv optErr **) + +let expressionless_genv ge = + let vars = ge.Genv.genv_vars in + let funcs = ge.Genv.genv_funcs in let methods = ge.Genv.genv_methods in + let defs = ge.Genv.genv_defs in let fundefs = ge.Genv.genv_fundefs in let methoddefs = ge.Genv.genv_methoddefs in let constructor = ge.Genv.genv_constructor in bind (Obj.magic coq_Monad_optErr) - (methodical_multiplexer_body methods methoddefs) (fun multiplexer_body -> - ret (Obj.magic coq_Monad_optErr) - (app (methodical_opt_function constructor) (Coq_cons ((Slabel body), - (Coq_cons ((Spush (Coq_inl (Vint sb))), (Coq_cons ((Spush (Coq_inl - (Vint sp))), (Coq_cons (Smstore, (Coq_cons ((Spush (Coq_inl (Vint - (Int256.repr (two_power_nat sg_shift))))), (Coq_cons ((Spush z0), - (Coq_cons (Scalldataload, (Coq_cons ((Sbinop (Odiv, Coq_false)), - (app multiplexer_body (Coq_cons (Srevert, - (app (methodical_methods methoddefs) (methodical_fundefs fundefs))))))))))))))))))))))) + (Obj.magic expressionless_fundefs fundefs) (fun fundefs0 -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic expressionless_methods methoddefs) (fun methoddefs0 -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic expressionless_constructor constructor) + (fun constructor0 -> + ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; + Genv.genv_funcs = funcs; Genv.genv_methods = methods; + Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; + Genv.genv_methoddefs = methoddefs0; Genv.genv_constructor = (Some + constructor0) }))) -(** val methodical_genv : program -> Language.genv optErr **) +(** val expressionless_program : + program -> StmtExpressionless.program optErr **) -let methodical_genv p = - bind (Obj.magic coq_Monad_optErr) (Obj.magic methodical_main p) - (fun main_code -> - let ge = fst p in - let body = snd p in - let vars = ge.Genv.genv_vars in - let defs = ge.Genv.genv_defs in - (match label_verify main_code with - | Coq_true -> - Success { genv_vars = vars; genv_defs = defs; genv_main = main_code; - genv_main_entrypoint = body } - | Coq_false -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) +let expressionless_program = function +| Coq_pair (ge, body) -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic expressionless_genv ge) + (fun cge -> + match label_verify cge with + | Coq_true -> ret (Obj.magic coq_Monad_optErr) (Coq_pair (cge, body)) + | Coq_false -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) diff --git a/src/backend/extraction/Gen8.mli b/src/backend/extraction/Gen8.mli index c905d26..cfbf553 100644 --- a/src/backend/extraction/Gen8.mli +++ b/src/backend/extraction/Gen8.mli @@ -1,35 +1,53 @@ open AST open Ascii -open Cop +open BinNums open Datatypes +open ExpStacked open Globalenvs open Integers -open Language -open List0 open LowValues open Maps0 +open MemoryModel open Monad open OptErrMonad -open Options -open Semantics3 -open StackEnv +open Semantics2 open StmtExpressionless +open StmtStacked open String0 -open Zpower +open Trees -val methodical_fundefs : coq_function PTree.t -> code +val expressionless_expr : expr -> StmtExpressionless.statement -val methodical_opt_function : coq_function option -> code +val pops : nat -> StmtExpressionless.statement list -val methodical_methods : coq_function option IntMap.t -> code +val cleanup : nat -> StmtExpressionless.statement list -val label_method_starts_with : coq_function -> label optErr +val expressionless_rt : ExpStacked.ret_type -> ret_type optErr -val sg_val : Int.int -> coq_val +val fetch_args : + function_kind -> nat -> nat -> StmtExpressionless.statement list -val methodical_multiplexer_body : - Int.int list -> coq_function option IntMap.t -> code optErr +val extract_lbl : typed_label -> label -val methodical_main : program -> code optErr +val expressionless_stm : + statement -> function_kind -> StmtExpressionless.statement list optErr -val methodical_genv : program -> Language.genv optErr +val expressionless_code : + code -> function_kind -> StmtExpressionless.code optErr + +val expressionless_function : + function_kind -> coq_function -> StmtExpressionless.coq_function optErr + +val expressionless_fundefs : + coq_function PTree.t -> StmtExpressionless.coq_function PTree.t optErr + +val expressionless_methods : + coq_function option IntMap.t -> StmtExpressionless.coq_function option + IntMap.t optErr + +val expressionless_constructor : + coq_function option -> StmtExpressionless.coq_function optErr + +val expressionless_genv : genv -> StmtExpressionless.genv optErr + +val expressionless_program : program -> StmtExpressionless.program optErr diff --git a/src/backend/extraction/Gen9.ml b/src/backend/extraction/Gen9.ml index b527e78..68aceb9 100644 --- a/src/backend/extraction/Gen9.ml +++ b/src/backend/extraction/Gen9.ml @@ -1,1913 +1,385 @@ open AST open Ascii -open BinNums -open BinPos open Cop -open Ctypes open Datatypes -open ExpCintptr -open GlobalenvCompile open Globalenvs -open Int0 open Integers -open Language0 +open Language open List0 +open LowValues open Maps0 open Monad -open Nat0 open OptErrMonad open Options -open PeanoNat +open Semantics3 open StackEnv -open StmtCintptr +open StmtExpressionless open String0 -open Structure -open Trees -open Values +open Zpower -let __ = let rec f _ = Obj.repr f in Obj.repr f +(** val methodical_fundefs : coq_function PTree.t -> code **) -(** val wasm_expr : - nat PTree.t -> ExpCintptr.expr -> bool -> instr list optErr **) +let methodical_fundefs t0 = + flat_map fn_code (map snd (PTree.elements t0)) -let rec wasm_expr temps e rvalue = - match e with - | Econst_int (_, _) -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, +(** val methodical_opt_function : coq_function option -> code **) + +let methodical_opt_function = function +| Some f -> fn_code f +| None -> Coq_nil + +(** val methodical_methods : coq_function option IntMap.t -> code **) + +let methodical_methods methods = + flat_map methodical_opt_function (map snd (PTree.elements (snd methods))) + +(** val label_method_starts_with : coq_function -> label optErr **) + +let label_method_starts_with m = + match fn_code m with + | Coq_nil -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Econst_int256 (i, _) -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Const_256 i), Coq_nil)) - | Etempvar (i, _) -> - (match rvalue with - | Coq_true -> - bind (Obj.magic coq_Monad_optErr) - (fromOption (PTree.get i (Obj.magic temps)) (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - (fun ind -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Local_get ind), - Coq_nil))) - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + | Coq_cons (s, _) -> + (match s with + | Slabel l -> ret (Obj.magic coq_Monad_optErr) l + | _ -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Esload (e0, _) -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e0 Coq_false) - (fun e' -> - ret (Obj.magic coq_Monad_optErr) - (app e' - (app - (set_eeiOffset Coq_nil Coq_pathOffset (S (S (S (S (S (S (S (S - O))))))))) - (app (Coq_cons ((Const_nat - (get_idx_eei_mem_offset Coq_pathOffset)), (Coq_cons ((Const_nat - (get_idx_eei_mem_offset Coq_resultOffset)), (Coq_cons ((Call - (get_idx_eei Coq_eei_storageLoad)), Coq_nil)))))) - (load_resultOffset (S O)))))) - | Emload (e0, t0) -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e0 Coq_false) - (fun e' -> - bind (Obj.magic coq_Monad_optErr) - (match t0 with - | Tpointer (p, _) -> - (match p with - | Coq_stor -> - Success (Coq_cons ((Const_nat (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S O))))))))))))))))))))))))))))))))), (Coq_cons ((Binop - (Coq_i32 IOp32.Add)), (Coq_cons ((Global_set - scratch_global_idx2), - (app Coq_nil - (load_len_hash (S (S (S (S (S (S (S (S O)))))))))))))))) - | _ -> Success (Coq_cons ((Memop (Coq_i32 IOp32.Load)), Coq_nil))) - | _ -> Success (Coq_cons ((Memop (Coq_i32 IOp32.Load)), Coq_nil))) - (fun l -> ret (Obj.magic coq_Monad_optErr) (app e' l))) - | Eaddr (e0, _) -> - (match rvalue with - | Coq_true -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e0 Coq_false) - (fun e' -> ret (Obj.magic coq_Monad_optErr) e') - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Eunop (o, e0, _) -> - (match rvalue with - | Coq_true -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e0 Coq_true) - (fun e' -> - bind (Obj.magic coq_Monad_optErr) - (match o with - | Onotbool -> Success (Coq_cons ((Testop (Coq_i32 __)), Coq_nil)) - | Onotint -> - Success (Coq_cons ((Call (get_idx_aux Coq_aux_notint)), - Coq_nil)) - | Oneg -> - Success (Coq_cons ((Const (Coq_i32 I32.zero)), (Coq_cons - ((Binop (Coq_i32 IOp32.Sub)), Coq_nil)))) - | Osha_1 -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - (fun o' -> ret (Obj.magic coq_Monad_optErr) (app e' o'))) - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Ebinop (o, e1, e2, _) -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e1 rvalue) (fun e1' -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e2 Coq_true) - (fun e2' -> - bind (Obj.magic coq_Monad_optErr) - (match o with - | Oadd -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Add)), Coq_nil)) - | Osub -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Sub)), Coq_nil)) - | Omul -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Mul)), Coq_nil)) - | Odiv -> - Success (Coq_cons ((Binop (Coq_i32 (IOp32.Div Unsigned))), - Coq_nil)) - | Omod -> - Success (Coq_cons ((Binop (Coq_i32 (IOp32.Rem Unsigned))), - Coq_nil)) - | Oexp -> - Success (Coq_cons ((Call (get_idx_aux Coq_aux_pow)), Coq_nil)) - | Oand -> Success (Coq_cons ((Binop (Coq_i32 IOp32.And)), Coq_nil)) - | Oor -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Or)), Coq_nil)) - | Oxor -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Xor)), Coq_nil)) - | Oshl -> Success (Coq_cons ((Binop (Coq_i32 IOp32.Shl)), Coq_nil)) - | Oshr -> - Success (Coq_cons ((Binop (Coq_i32 (IOp32.Shr Unsigned))), - Coq_nil)) - | Oeq -> Success (Coq_cons ((Relop (Coq_i32 IOp32.Eq)), Coq_nil)) - | One -> Success (Coq_cons ((Relop (Coq_i32 IOp32.Ne)), Coq_nil)) - | Olt -> - Success (Coq_cons ((Relop (Coq_i32 (IOp32.Lt Unsigned))), - Coq_nil)) - | Ogt -> - Success (Coq_cons ((Relop (Coq_i32 (IOp32.Gt Unsigned))), - Coq_nil)) - | Ole -> - Success (Coq_cons ((Relop (Coq_i32 (IOp32.Le Unsigned))), - Coq_nil)) - | Oge -> - Success (Coq_cons ((Relop (Coq_i32 (IOp32.Ge Unsigned))), - Coq_nil)) - | Osha_2 -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - (fun o' -> ret (Obj.magic coq_Monad_optErr) (app e1' (app e2' o'))))) - | Ecall0 (b, _) -> - (match rvalue with - | Coq_true -> - bind (Obj.magic coq_Monad_optErr) (wasm_builtin0 b) (fun b' -> - ret (Obj.magic coq_Monad_optErr) b') - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Ecall1 (b, e0, _) -> - (match rvalue with - | Coq_true -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e0 Coq_true) - (fun e' -> - bind (Obj.magic coq_Monad_optErr) (Success (wasm_builtin1 b)) - (fun b' -> ret (Obj.magic coq_Monad_optErr) (app e' b'))) - | Coq_false -> - Error (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - -(** val wasm_exprs : - nat PTree.t -> ExpCintptr.expr list -> instr list optErr **) - -let rec wasm_exprs temps = function -| Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil -| Coq_cons (e, rest) -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps e Coq_true) (fun e' -> - bind (Obj.magic coq_Monad_optErr) (wasm_exprs temps rest) (fun rest' -> - ret (Obj.magic coq_Monad_optErr) (app rest' e'))) + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -(** val optident : nat PTree.t -> ident option -> instrs optErr **) +(** val sg_val : Int.int -> coq_val **) -let optident temps = function -| Some b -> - (match PTree.get b temps with - | Some v -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Local_set v), Coq_nil)) - | None -> - Error (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))) -| None -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (Drop, Coq_nil)) +let sg_val sg = + Vint (Int256.repr (Int.unsigned sg)) -(** val wasm_statement : - (positive, nat) prod list -> nat PTree.t -> coq_type -> statement -> nat - -> instrs optErr **) +(** val methodical_multiplexer_body : + Int.int list -> coq_function option IntMap.t -> code optErr **) -let rec wasm_statement funident_map temps return_type s cstack_depth = - match s with - | Sskip -> ret (Obj.magic coq_Monad_optErr) (Coq_cons (Nop, Coq_nil)) - | Ssassign (lv, rv) -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps lv Coq_false) - (fun lv' -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps rv Coq_true) - (fun rv' -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app - (set_eeiOffset Coq_nil Coq_pathOffset (S (S (S (S (S (S (S (S - O))))))))) - (app (set_eeiOffset rv' Coq_valueOffset (S O)) (Coq_cons - ((Const_nat (get_idx_eei_mem_offset Coq_pathOffset)), - (Coq_cons ((Const_nat - (get_idx_eei_mem_offset Coq_valueOffset)), (Coq_cons ((Call - (get_idx_eei Coq_eei_storageStore)), Coq_nil))))))))))) - | Smassign (lv, rv) -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps lv Coq_false) - (fun lv' -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps rv Coq_true) - (fun rv' -> - match typeof lv with - | Tpointer (_, _) -> - (match rv with - | Econst_int (_, _) -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) - | Econst_int256 (_, _) -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) - | Etempvar (_, _) -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) - | Esload (_, _) -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) - | Emload (_, t0) -> - (match t0 with - | Tpointer (mem, _) -> - (match mem with - | Coq_mem -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), - Coq_nil)))) - | Coq_stor -> - ret (Obj.magic coq_Monad_optErr) - (app rv' - (store_len lv' (S (S (S (S (S (S (S (S O)))))))))) - | Coq_call -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), - Coq_nil))))) - | _ -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Eaddr (_, _) -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) - | Eunop (_, _, _) -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) - | Ebinop (_, _, _, _) -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) - | Ecall0 (_, _) -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))) - | Ecall1 (_, _, _) -> - ret (Obj.magic coq_Monad_optErr) - (app lv' - (app rv' (Coq_cons ((Memop (Coq_i32 IOp32.Store)), Coq_nil))))) - | _ -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Sset (id, rv) -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps rv Coq_true) - (fun rv' -> - bind (Obj.magic coq_Monad_optErr) - (match rv with - | Emload (_, t0) -> - (match t0 with - | Tpointer (p, _) -> - (match p with - | Coq_stor -> - Success (Coq_cons (Drop, (Coq_cons (Drop, (Coq_cons (Drop, - (Coq_cons (Drop, (Coq_cons (Drop, (Coq_cons (Drop, - (Coq_cons (Drop, Coq_nil)))))))))))))) - | _ -> Success Coq_nil) - | _ -> Success Coq_nil) - | _ -> Success Coq_nil) (fun d -> - let set_lookup = - match PTree.get id temps with - | Some v -> Coq_cons ((Local_set v), Coq_nil) - | None -> Coq_nil - in - ret (Obj.magic coq_Monad_optErr) (app rv' (app d set_lookup)))) - | Scall (retval, lab, args) -> - bind (Obj.magic coq_Monad_optErr) (optident temps retval) (fun t0 -> - bind (Obj.magic coq_Monad_optErr) (wasm_exprs temps args) (fun args' -> - let ido = - fold_left (fun xs x -> - match xs with - | Some _ -> xs - | None -> - (match Pos.eqb (fst x) lab with - | Coq_true -> Some (snd x) - | Coq_false -> None)) funident_map None - in - (match ido with - | Some id -> - ret (Obj.magic coq_Monad_optErr) - (app args' (app (Coq_cons ((Call id), Coq_nil)) t0)) - | None -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Ssequence (s1, s2) -> - bind (Obj.magic coq_Monad_optErr) - (wasm_statement funident_map temps return_type s1 cstack_depth) - (fun x1 -> - bind (Obj.magic coq_Monad_optErr) - (wasm_statement funident_map temps return_type s2 cstack_depth) - (fun x2 -> ret (Obj.magic coq_Monad_optErr) (app x1 x2))) - | Sifthenelse (c, strue, sfalse) -> +let rec methodical_multiplexer_body methods methoddefs = + match methods with + | Coq_nil -> ret (Obj.magic coq_Monad_optErr) Coq_nil + | Coq_cons (sg, rest) -> bind (Obj.magic coq_Monad_optErr) - (wasm_statement funident_map temps return_type sfalse - (add cstack_depth (S O))) (fun nfalse -> + (methodical_multiplexer_body rest methoddefs) (fun rest' -> bind (Obj.magic coq_Monad_optErr) - (wasm_statement funident_map temps return_type strue - (add cstack_depth (S O))) (fun ntrue -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps c Coq_true) - (fun cond -> - ret (Obj.magic coq_Monad_optErr) - (app cond (Coq_cons ((If ((BT_valtype None), ntrue, nfalse)), - Coq_nil)))))) - | Sloop sbody -> - bind (Obj.magic coq_Monad_optErr) - (wasm_statement funident_map temps return_type sbody - (add cstack_depth (S O))) (fun n2 -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Loop ((BT_valtype None), - n2)), Coq_nil))) - | Sbreak -> - (match Nat.eqb cstack_depth O with + (fromOption (IntMap.get sg (Obj.magic methoddefs)) (String ((Ascii + (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii + (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, + Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), + EmptyString))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) + (fun m -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic label_method_starts_with m) (fun l -> + ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Sdup O), (Coq_cons + ((Spush (Coq_inl (sg_val sg))), (Coq_cons ((Sbinop (Oeq, + Coq_false)), (Coq_cons ((Spush (Coq_inr l)), (Coq_cons (Sjumpi, + rest'))))))))))))) + +(** val methodical_main : program -> code optErr **) + +let methodical_main p = + let ge = fst p in + let body = snd p in + let methods = ge.Genv.genv_methods in + let fundefs = ge.Genv.genv_fundefs in + let methoddefs = ge.Genv.genv_methoddefs in + let constructor = ge.Genv.genv_constructor in + bind (Obj.magic coq_Monad_optErr) + (methodical_multiplexer_body methods methoddefs) (fun multiplexer_body -> + ret (Obj.magic coq_Monad_optErr) + (app (methodical_opt_function constructor) (Coq_cons ((Slabel body), + (Coq_cons ((Spush (Coq_inl (Vint sb))), (Coq_cons ((Spush (Coq_inl + (Vint sp))), (Coq_cons (Smstore, (Coq_cons ((Spush (Coq_inl (Vint + (Int256.repr (two_power_nat sg_shift))))), (Coq_cons ((Spush z0), + (Coq_cons (Scalldataload, (Coq_cons ((Sbinop (Odiv, Coq_false)), + (app multiplexer_body (Coq_cons (Srevert, + (app (methodical_methods methoddefs) (methodical_fundefs fundefs))))))))))))))))))))))) + +(** val methodical_genv : program -> Language.genv optErr **) + +let methodical_genv p = + bind (Obj.magic coq_Monad_optErr) (Obj.magic methodical_main p) + (fun main_code -> + let ge = fst p in + let body = snd p in + let vars = ge.Genv.genv_vars in + let defs = ge.Genv.genv_defs in + (match label_verify main_code with | Coq_true -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, + Success { genv_vars = vars; genv_defs = defs; genv_main = main_code; + genv_main_entrypoint = body } + | Coq_false -> + Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), + (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, + ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - | Coq_false -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Br cstack_depth), - Coq_nil))) - | Sreturn retval -> - bind (Obj.magic coq_Monad_optErr) - (match retval with - | Some i -> wasm_expr temps (Etempvar (i, return_type)) Coq_true - | None -> Success (Coq_cons ((Const_nat O), Coq_nil))) (fun rv' -> - ret (Obj.magic coq_Monad_optErr) (app rv' (Coq_cons (Return, Coq_nil)))) - | Shash (ex1, ex2, exo) -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps ex1 Coq_true) - (fun ex1' -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps ex2 Coq_true) - (fun ex2' -> - bind (Obj.magic coq_Monad_optErr) - (match exo with - | Some ex -> wasm_expr temps ex Coq_true - | None -> ret (Obj.magic coq_Monad_optErr) ex2') (fun exo' -> - ret (Obj.magic coq_Monad_optErr) - (app ex2' - (app (Coq_cons ((Memop (Coq_i32 IOp32.Load)), Coq_nil)) - (app exo' - (app (Coq_cons ((Memop (Coq_i32 IOp32.Load)), Coq_nil)) - (app callExternalSha256TwoParam - (store_len ex1' (S (S (S (S (S (S (S (S O))))))))))))))))) - | Stransfer (a, v) -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps a Coq_true) (fun a' -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps v Coq_true) - (fun v' -> - ret (Obj.magic coq_Monad_optErr) - (app (set_eeiOffset a' Coq_addressOffset (S O)) - (app (set_eeiOffset v' Coq_valueOffset (S O)) - (app (set_eeiOffset Coq_nil Coq_dataOffset (S O)) (Coq_cons - ((Const_nat default_gas_limit), (Coq_cons ((Cvtop (Coq_i32 - (IOp32.Extend_i32 Unsigned))), (Coq_cons ((Const_nat - (get_idx_eei_mem_offset Coq_addressOffset)), (Coq_cons - ((Const_nat (get_idx_eei_mem_offset Coq_valueOffset)), - (Coq_cons ((Const_nat - (get_idx_eei_mem_offset Coq_dataOffset)), (Coq_cons - ((Const_nat O), (Coq_cons ((Call (get_idx_eei Coq_eei_call)), - Coq_nil))))))))))))))))))) - | Scallmethod (a, _, _, v, args) -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps a Coq_true) (fun a' -> - bind (Obj.magic coq_Monad_optErr) - (fold_left (fun xs x -> - bind (Obj.magic coq_Monad_optErr) xs (fun xsl -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps x Coq_true) - (fun rst -> ret (Obj.magic coq_Monad_optErr) (app xsl rst)))) - args (ret (Obj.magic coq_Monad_optErr) Coq_nil)) (fun _ -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps v Coq_true) - (fun v' -> - ret (Obj.magic coq_Monad_optErr) - (app (set_eeiOffset a' Coq_addressOffset (S O)) - (app (set_eeiOffset v' Coq_valueOffset (S O)) - (app (set_eeiOffset Coq_nil Coq_dataOffset (S O)) (Coq_cons - ((Const_nat default_gas_limit), (Coq_cons ((Cvtop (Coq_i32 - (IOp32.Extend_i32 Unsigned))), (Coq_cons ((Const_nat - (get_idx_eei_mem_offset Coq_addressOffset)), (Coq_cons - ((Const_nat (get_idx_eei_mem_offset Coq_valueOffset)), - (Coq_cons ((Const_nat - (get_idx_eei_mem_offset Coq_dataOffset)), (Coq_cons - ((Const_nat O), (Coq_cons ((Call - (get_idx_eei Coq_eei_call)), Coq_nil)))))))))))))))))))) - | Slog (topics, args) -> - bind (Obj.magic coq_Monad_optErr) - (fold_left (fun xs x -> - bind (Obj.magic coq_Monad_optErr) xs (fun xsl -> - bind (Obj.magic coq_Monad_optErr) - (match x with - | Econst_int256 (i, _) -> - Success (Coq_cons ((Const_256 - (Int256.shru - (Int256.shl i - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO (Coq_xI (Coq_xI Coq_xH)))))))))) - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 - (Int256.shru - (Int256.shl i - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO (Coq_xO (Coq_xI Coq_xH)))))))))) - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 - (Int256.shru - (Int256.shl i - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO (Coq_xI (Coq_xO Coq_xH)))))))))) - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 - (Int256.shru - (Int256.shl i - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO (Coq_xO (Coq_xO Coq_xH)))))))))) - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 - (Int256.shru - (Int256.shl i - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO (Coq_xI Coq_xH))))))))) - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 - (Int256.shru - (Int256.shl i - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO (Coq_xO Coq_xH))))))))) - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 - (Int256.shru - (Int256.shl i - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xO Coq_xH)))))))) - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xI (Coq_xI Coq_xH))))))))))), (Coq_cons ((Const_256 - (Int256.shru (Int256.shl i (Int256.repr Z0)) - (Int256.repr (Zpos (Coq_xO (Coq_xO (Coq_xO (Coq_xO (Coq_xO - (Coq_xI (Coq_xI Coq_xH))))))))))), - Coq_nil)))))))))))))))) - | _ -> - Error (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), - EmptyString))))))))))))))))))))))))))))))))))))))))))))) - (fun rst -> ret (Obj.magic coq_Monad_optErr) (app xsl rst)))) - topics (ret (Obj.magic coq_Monad_optErr) Coq_nil)) (fun topics' -> - bind (Obj.magic coq_Monad_optErr) - (fold_left (fun xs x -> - bind (Obj.magic coq_Monad_optErr) xs (fun xsl -> - bind (Obj.magic coq_Monad_optErr) (wasm_expr temps x Coq_true) - (fun rst -> ret (Obj.magic coq_Monad_optErr) (app xsl rst)))) - args (ret (Obj.magic coq_Monad_optErr) Coq_nil)) (fun args' -> - let topicsnumber = - match topics with - | Coq_nil -> O - | Coq_cons (_, l) -> - (match l with - | Coq_nil -> S O - | Coq_cons (_, l0) -> - (match l0 with - | Coq_nil -> S (S O) - | Coq_cons (_, l1) -> - (match l1 with - | Coq_nil -> S (S (S O)) - | Coq_cons (_, l2) -> - (match l2 with - | Coq_nil -> S (S (S (S O))) - | Coq_cons (_, _) -> S (S (S (S (S O)))))))) - in - let argsnumber = length args in - let compiledInstrs = - app args' - (app (storeLogArgs argsnumber) - (app topics' - (app - (match topics with - | Coq_nil -> Coq_nil - | Coq_cons (_, l) -> - (match l with - | Coq_nil -> - set_eeiOffset Coq_nil Coq_topic1Offset (S (S (S (S (S - (S (S (S O)))))))) - | Coq_cons (_, l0) -> - (match l0 with - | Coq_nil -> - app - (set_eeiOffset Coq_nil Coq_topic1Offset (S (S (S - (S (S (S (S (S O))))))))) - (set_eeiOffset Coq_nil Coq_topic2Offset (S (S (S - (S (S (S (S (S O))))))))) - | Coq_cons (_, l1) -> - (match l1 with - | Coq_nil -> - app - (set_eeiOffset Coq_nil Coq_topic1Offset (S (S - (S (S (S (S (S (S O))))))))) - (app - (set_eeiOffset Coq_nil Coq_topic2Offset (S - (S (S (S (S (S (S (S O))))))))) - (set_eeiOffset Coq_nil Coq_topic3Offset (S - (S (S (S (S (S (S (S O)))))))))) - | Coq_cons (_, l2) -> - (match l2 with - | Coq_nil -> - app - (set_eeiOffset Coq_nil Coq_topic1Offset (S - (S (S (S (S (S (S (S O))))))))) - (app - (set_eeiOffset Coq_nil Coq_topic2Offset - (S (S (S (S (S (S (S (S O))))))))) - (app - (set_eeiOffset Coq_nil - Coq_topic3Offset (S (S (S (S (S (S - (S (S O))))))))) - (set_eeiOffset Coq_nil - Coq_topic4Offset (S (S (S (S (S (S - (S (S O))))))))))) - | Coq_cons (_, _) -> Coq_nil))))) (Coq_cons - ((Const_nat (get_idx_eei_mem_offset Coq_dataOffset)), - (Coq_cons ((Const_nat (mul argsnumber (S (S (S (S O)))))), - (Coq_cons ((Const_nat topicsnumber), (Coq_cons ((Const_nat - (get_idx_eei_mem_offset Coq_topic1Offset)), (Coq_cons - ((Const_nat (get_idx_eei_mem_offset Coq_topic2Offset)), - (Coq_cons ((Const_nat - (get_idx_eei_mem_offset Coq_topic3Offset)), (Coq_cons - ((Const_nat (get_idx_eei_mem_offset Coq_topic4Offset)), - (Coq_cons ((Call (get_idx_eei Coq_eei_log)), - Coq_nil))))))))))))))))))) - in - (match Nat.eqb topicsnumber (S (S (S (S (S O))))) with - | Coq_true -> - Error (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))) - | Coq_false -> ret (Obj.magic coq_Monad_optErr) compiledInstrs))) - | Srevert -> - ret (Obj.magic coq_Monad_optErr) (Coq_cons ((Const_nat - (get_idx_eei_mem_offset Coq_dataOffset)), (Coq_cons ((Const_nat - eei_revert_output_data_length), (Coq_cons ((Call - (get_idx_eei Coq_eei_revert)), Coq_nil)))))) - -(** val allocate_temps : (ident, coq_type) prod list -> nat -> nat PTree.t **) - -let rec allocate_temps ids base = - match ids with - | Coq_nil -> PTree.empty - | Coq_cons (p, rest) -> - let Coq_pair (id, _) = p in - PTree.set id base (allocate_temps rest (S base)) - -(** val allocate_all_temps : (ident, coq_type) prod list -> nat PTree.t **) - -let allocate_all_temps ids = - allocate_temps ids O - -(** val allocate_fn_temps : coq_function -> nat PTree.t **) - -let allocate_fn_temps f = - allocate_all_temps (app f.fn_params f.fn_temps) - -(** val wasm_function : - coq_Z PTree.t -> (positive, nat) prod list -> coq_function -> - Language0.coq_function optErr **) - -let wasm_function _ funident_map f = - let wasm_fb = - wasm_statement funident_map (allocate_fn_temps f) f.fn_return f.fn_body O - in - (match wasm_fb with - | Success insts -> - ret (Obj.magic coq_Monad_optErr) { Language0.fn_return = f.fn_return; - Language0.fn_params = f.fn_params; Language0.fn_temps = f.fn_temps; - Language0.fn_locals = f.fn_locals; Language0.fn_body = - (app (Coq_cons ((Const_256 sp), (Coq_cons ((Const_256 sb), (Coq_cons - ((Memop (Coq_i32 IOp32.Store)), Coq_nil)))))) insts) } - | Error msg -> - Error - (append (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_true, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_false, Coq_false, - Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii - (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), - (String ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, + Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_false, + Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String + ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, - Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_false, Coq_false, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + ((Ascii (Coq_true, Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_false, Coq_true, Coq_false, Coq_true, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, Coq_true, Coq_true, - Coq_false, Coq_true, Coq_false, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_false, Coq_false, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_false, Coq_true, Coq_false)), (String + ((Ascii (Coq_false, Coq_true, Coq_true, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_true, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_true, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_true, - Coq_true, Coq_true, Coq_false, Coq_false, Coq_false, Coq_true, - Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_true, - Coq_false, Coq_false, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_true, Coq_true, Coq_false, - Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, - Coq_true, Coq_true, Coq_true, Coq_false, Coq_true, Coq_false, - Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_true, + Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_false, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, Coq_true, Coq_false)), (String - ((Ascii (Coq_false, Coq_true, Coq_false, Coq_true, Coq_true, - Coq_true, Coq_false, Coq_false)), (String ((Ascii (Coq_false, - Coq_false, Coq_false, Coq_false, Coq_false, Coq_true, Coq_false, - Coq_false)), - EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - msg)) - -(** val wasm_constructor : - coq_function option -> coq_Z PTree.t -> (positive, nat) prod list -> - Language0.coq_function option optErr **) - -let wasm_constructor f global funident_map = - match f with - | Some f0 -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic wasm_function global funident_map f0) (fun cf -> - ret (Obj.magic coq_Monad_optErr) (Some cf)) - | None -> ret (Obj.magic coq_Monad_optErr) None - -(** val wasm_functions : - coq_function PTree.t -> coq_Z PTree.t -> (positive, nat) prod list -> - Language0.coq_function PTree.t optErr **) - -let wasm_functions defs global funident_map = - transl_tree (wasm_function global funident_map) defs - -(** val wasm_methoddefs : - coq_function option IntMap.t -> coq_Z PTree.t -> (positive, nat) prod - list -> Language0.coq_function option IntMap.t optErr **) - -let wasm_methoddefs defs global funident_map = - transl_map (wasm_function global funident_map) defs - -(** val wasm_func_identifier : - (positive, 'a1) prod list -> nat -> nat -> (positive, nat) prod list **) - -let rec wasm_func_identifier a idx base = - match a with - | Coq_nil -> Coq_nil - | Coq_cons (x, xs) -> - Coq_cons ((Coq_pair ((fst x), (add idx base))), - (wasm_func_identifier xs (add idx (S O)) base)) - -(** val wasm_genv : genv -> Language0.genv optErr **) - -let wasm_genv ge = - let vars = ge.Genv.genv_vars in - let names = ge.Genv.genv_funcs in - let fundefs = ge.Genv.genv_fundefs in - let sigs = ge.Genv.genv_methods in - let defs = ge.Genv.genv_defs in - let methoddefs = ge.Genv.genv_methoddefs in - let constructor = ge.Genv.genv_constructor in - let wasm_func_id_base = - add (length sigs) (match constructor with - | Some _ -> S O - | None -> O) - in - let funident_map = - wasm_func_identifier (PTree.elements fundefs) O - (add wasm_func_id_base (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S - (S (S (S (S (S (S (S (S (S (S - O))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) - in - bind (Obj.magic coq_Monad_optErr) - (Obj.magic allocate_addrs ge.Genv.genv_vars Z0 ge.Genv.genv_defs) - (fun allocated -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic wasm_functions fundefs allocated funident_map) - (fun fundefs0 -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic wasm_methoddefs methoddefs allocated funident_map) - (fun methoddefs0 -> - bind (Obj.magic coq_Monad_optErr) - (Obj.magic wasm_constructor constructor allocated funident_map) - (fun constructor0 -> - ret (Obj.magic coq_Monad_optErr) { Genv.genv_vars = vars; - Genv.genv_funcs = names; Genv.genv_methods = sigs; - Genv.genv_defs = defs; Genv.genv_fundefs = fundefs0; - Genv.genv_methoddefs = methoddefs0; Genv.genv_constructor = - constructor0 })))) + ((Ascii (Coq_true, Coq_false, Coq_false, Coq_true, Coq_false, + Coq_true, Coq_true, Coq_false)), (String ((Ascii (Coq_false, + Coq_true, Coq_true, Coq_false, Coq_false, Coq_true, Coq_true, + Coq_false)), (String ((Ascii (Coq_true, Coq_false, Coq_false, + Coq_true, Coq_true, Coq_true, Coq_true, Coq_false)), + EmptyString)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) diff --git a/src/backend/extraction/Gen9.mli b/src/backend/extraction/Gen9.mli index b9dd566..c905d26 100644 --- a/src/backend/extraction/Gen9.mli +++ b/src/backend/extraction/Gen9.mli @@ -1,63 +1,35 @@ open AST open Ascii -open BinNums -open BinPos open Cop -open Ctypes open Datatypes -open ExpCintptr -open GlobalenvCompile open Globalenvs -open Int0 open Integers -open Language0 +open Language open List0 +open LowValues open Maps0 open Monad -open Nat0 open OptErrMonad open Options -open PeanoNat +open Semantics3 open StackEnv -open StmtCintptr +open StmtExpressionless open String0 -open Structure -open Trees -open Values +open Zpower -val wasm_expr : nat PTree.t -> ExpCintptr.expr -> bool -> instr list optErr +val methodical_fundefs : coq_function PTree.t -> code -val wasm_exprs : nat PTree.t -> ExpCintptr.expr list -> instr list optErr +val methodical_opt_function : coq_function option -> code -val optident : nat PTree.t -> ident option -> instrs optErr +val methodical_methods : coq_function option IntMap.t -> code -val wasm_statement : - (positive, nat) prod list -> nat PTree.t -> coq_type -> statement -> nat -> - instrs optErr +val label_method_starts_with : coq_function -> label optErr -val allocate_temps : (ident, coq_type) prod list -> nat -> nat PTree.t +val sg_val : Int.int -> coq_val -val allocate_all_temps : (ident, coq_type) prod list -> nat PTree.t +val methodical_multiplexer_body : + Int.int list -> coq_function option IntMap.t -> code optErr -val allocate_fn_temps : coq_function -> nat PTree.t +val methodical_main : program -> code optErr -val wasm_function : - coq_Z PTree.t -> (positive, nat) prod list -> coq_function -> - Language0.coq_function optErr - -val wasm_constructor : - coq_function option -> coq_Z PTree.t -> (positive, nat) prod list -> - Language0.coq_function option optErr - -val wasm_functions : - coq_function PTree.t -> coq_Z PTree.t -> (positive, nat) prod list -> - Language0.coq_function PTree.t optErr - -val wasm_methoddefs : - coq_function option IntMap.t -> coq_Z PTree.t -> (positive, nat) prod list - -> Language0.coq_function option IntMap.t optErr - -val wasm_func_identifier : - (positive, 'a1) prod list -> nat -> nat -> (positive, nat) prod list - -val wasm_genv : genv -> Language0.genv optErr +val methodical_genv : program -> Language.genv optErr diff --git a/src/backend/extraction/Glue.ml b/src/backend/extraction/Glue.ml index 991ab16..2047ed4 100644 --- a/src/backend/extraction/Glue.ml +++ b/src/backend/extraction/Glue.ml @@ -2,17 +2,18 @@ open AST open BinNums open Datatypes open EVM -open Gen2 +open Gen3 open Gen1 +open Gen6 open Gen5 -open Gen4 open Gen -open Gen3 +open Gen4 open Gen0 -open Gen7 +open Gen2 open Gen8 -open Gen6 open Gen9 +open Gen7 +open Gen10 open GlobalenvCompile open Integers open Monad @@ -44,7 +45,9 @@ let optm_genv ge = bind (Obj.magic coq_Monad_optErr) (Obj.magic clocal_genv clike) (fun clocal -> bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_genv clocal) - cgraph_viz)) + (fun cintptr -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cgraph_genv cintptr) + cgraph_viz))) (** val optm_clash : genv -> (positive, positive list) prod list list optErr **) @@ -54,7 +57,9 @@ let optm_clash ge = bind (Obj.magic coq_Monad_optErr) (Obj.magic clocal_genv clike) (fun clocal -> bind (Obj.magic coq_Monad_optErr) (Obj.magic cintptr_genv clocal) - clash_viz)) + (fun cintptr -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cgraph_genv cintptr) + clash_viz))) (** val full_compile_genv : genv -> (evm list, label) prod optErr **) @@ -66,24 +71,27 @@ let full_compile_genv ge = (fun cintptr -> bind (Obj.magic coq_Monad_optErr) (Obj.magic cgraph_genv cintptr) (fun cgraph -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic cbasic_genv cgraph) - (fun cbasic -> - bind (Obj.magic coq_Monad_optErr) (Obj.magic clinear_genv cbasic) - (fun clinear -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic copt_genv cgraph) + (fun copt -> + bind (Obj.magic coq_Monad_optErr) (Obj.magic cbasic_genv copt) + (fun cbasic -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic clabeled_program clinear) (fun clabeled -> + (Obj.magic clinear_genv cbasic) (fun clinear -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic stacked_program clabeled) (fun stacked -> + (Obj.magic clabeled_program clinear) (fun clabeled -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic expressionless_program stacked) - (fun expressionless -> + (Obj.magic stacked_program clabeled) (fun stacked -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic methodical_genv expressionless) - (fun methodical -> + (Obj.magic expressionless_program stacked) + (fun expressionless -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic genv_compiled methodical) (fun program -> + (Obj.magic methodical_genv expressionless) + (fun methodical -> bind (Obj.magic coq_Monad_optErr) - (Obj.magic get_main_entrypoint methodical) - (fun main_entrypoint -> - ret (Obj.magic coq_Monad_optErr) (Coq_pair - (program, main_entrypoint)))))))))))))) + (Obj.magic genv_compiled methodical) + (fun program -> + bind (Obj.magic coq_Monad_optErr) + (Obj.magic get_main_entrypoint methodical) + (fun main_entrypoint -> + ret (Obj.magic coq_Monad_optErr) (Coq_pair + (program, main_entrypoint))))))))))))))) diff --git a/src/backend/extraction/Glue.mli b/src/backend/extraction/Glue.mli index 2ed6c8f..c4a4097 100644 --- a/src/backend/extraction/Glue.mli +++ b/src/backend/extraction/Glue.mli @@ -2,17 +2,18 @@ open AST open BinNums open Datatypes open EVM -open Gen2 +open Gen3 open Gen1 +open Gen6 open Gen5 -open Gen4 open Gen -open Gen3 +open Gen4 open Gen0 -open Gen7 +open Gen2 open Gen8 -open Gen6 open Gen9 +open Gen7 +open Gen10 open GlobalenvCompile open Integers open Monad diff --git a/src/backend/extraction/LanguageExt.ml b/src/backend/extraction/LanguageExt.ml index 3bc27b9..2869d4d 100755 --- a/src/backend/extraction/LanguageExt.ml +++ b/src/backend/extraction/LanguageExt.ml @@ -102,6 +102,10 @@ let rec show_expr show_tmp expr = else expr' +let show_expr_opt show_tmp = function + | D.None -> "" + | D.Some e -> show_expr show_tmp e + (* val show_stmt : stmt -> string list *) let rec show_stmt show_tmp indent stmt = let show_stmt = show_stmt show_tmp in @@ -122,12 +126,12 @@ let rec show_stmt show_tmp indent stmt = List.append (show_stmt false s1) (show_stmt false s2) | Sifthenelse (e, s1, s2) -> pad (List.flatten [ - [ sprintf "if (%s) {" (show_expr e) ]; - show_stmt true s1; - [ "} else {" ]; - show_stmt true s2; - [ "}" ]; - ]) + [ sprintf "if (%s) {" (show_expr e) ]; + show_stmt true s1; + [ "} else {" ]; + show_stmt true s2; + [ "}" ]; + ]) | Sloop s -> pad (List.flatten [ (* TODO: add "true" and "false" keywords *) @@ -144,8 +148,16 @@ let rec show_stmt show_tmp indent stmt = [ sprintf "return %s;" (show_tmp i) ]) | Stransfer (e1, e2) -> [ sprintf "transfer(%s, %s);" (show_expr e1) (show_expr e2) ] - | Scallmethod (e1, idents, v, e2, es) -> - [ "// CALLMETHOD unimplemented" ] + | Scallmethod (e1, idents, e2, v, g, es) -> + [ sprintf "callmethod(%s);" (String.concat "; " [ + (show_expr e1); + (String.concat ", " (List.map show_tmp (DE.caml_list idents))); + (show_coq_int e2); + (show_expr v); + (show_expr_opt show_tmp g); + (String.concat ", " (List.map show_expr (DE.caml_list es))) + ]) + ] | Slog (topics, args) -> let f l = String.concat ", " (List.map show_expr l) in [ sprintf "emit(%s; %s);" (f (DE.caml_list topics)) (f (DE.caml_list args)) ] | Srevert -> diff --git a/src/backend/extraction/Semantics2.ml b/src/backend/extraction/Semantics2.ml index 911cb1e..450ca9d 100644 --- a/src/backend/extraction/Semantics2.ml +++ b/src/backend/extraction/Semantics2.ml @@ -449,9 +449,9 @@ let stm_eq_dec stm stm' = | Stransfer -> (match stm' with | Stransfer -> Coq_left | _ -> Coq_right) - | Scallmethod (x, x0, x1) -> + | Scallargs (x, x0, x1) -> (match stm' with - | Scallmethod (i0, n1, n2) -> + | Scallargs (i0, n1, n2) -> (match Int.eq_dec x i0 with | Coq_left -> (match Nat.eq_dec x0 n1 with @@ -459,6 +459,19 @@ let stm_eq_dec stm stm' = | Coq_right -> Coq_right) | Coq_right -> Coq_right) | _ -> Coq_right) + | Scallmethod x -> + (match stm' with + | Scallmethod b0 -> + (match x with + | Coq_true -> + (match b0 with + | Coq_true -> Coq_left + | Coq_false -> Coq_right) + | Coq_false -> + (match b0 with + | Coq_true -> Coq_right + | Coq_false -> Coq_left)) + | _ -> Coq_right) | Slog (x, x0) -> (match stm' with | Slog (n1, n2) -> diff --git a/src/backend/extraction/StmCompile.ml b/src/backend/extraction/StmCompile.ml index dee84a8..83ee4d4 100644 --- a/src/backend/extraction/StmCompile.ml +++ b/src/backend/extraction/StmCompile.ml @@ -217,8 +217,12 @@ let stm_compiled stm _ code_label = | Sjumpi -> command_compiled Coq_evm_jumpi | Shash -> command_compiled Coq_evm_sha3 | Stransfer -> command_compiled Coq_evm_call - | Scallmethod (sg, args, rvcount) -> - append_compiled Coq_evm_call (push_public_args args O rvcount sg) + | Scallargs (sg, args, rvcount) -> push_public_args args O rvcount sg + | Scallmethod b -> + (match b with + | Coq_true -> command_compiled Coq_evm_call + | Coq_false -> + append_compiled Coq_evm_call (command_compiled Coq_evm_gas)) | Slog (ntopics, nargs) -> (match Nat.ltb nargs (S (S (S (S (S O))))) with | Coq_true -> diff --git a/src/backend/extraction/StmtCGraph.ml b/src/backend/extraction/StmtCGraph.ml index 5b155d7..f572f89 100644 --- a/src/backend/extraction/StmtCGraph.ml +++ b/src/backend/extraction/StmtCGraph.ml @@ -19,7 +19,8 @@ type statement = | Sdone | Shash of expr * expr * expr option * node | Stransfer of expr * expr * node * node -| Scallmethod of expr * ident list * Int.int * expr * expr list * node * node +| Scallmethod of expr * ident list * Int.int * expr * expr option * expr list + * node * node | Slog of expr list * expr list * node | Srevert @@ -29,7 +30,8 @@ type coq_function = { fn_return : coq_type; fn_params : (ident, coq_type) prod list; fn_temps : (ident, coq_type) prod list; fn_locals : (ident, coq_type) prod list; - fn_code : code; fn_entrypoint : node } + fn_code : code; fn_entrypoint : node; fn_nextnode : + node } (** val fn_return : coq_function -> coq_type **) @@ -55,4 +57,8 @@ let fn_code x = x.fn_code let fn_entrypoint x = x.fn_entrypoint +(** val fn_nextnode : coq_function -> node **) + +let fn_nextnode x = x.fn_nextnode + type genv = (coq_function, coq_type) Genv.t diff --git a/src/backend/extraction/StmtCGraph.mli b/src/backend/extraction/StmtCGraph.mli index 34679ea..7495bc9 100644 --- a/src/backend/extraction/StmtCGraph.mli +++ b/src/backend/extraction/StmtCGraph.mli @@ -19,7 +19,8 @@ type statement = | Sdone | Shash of expr * expr * expr option * node | Stransfer of expr * expr * node * node -| Scallmethod of expr * ident list * Int.int * expr * expr list * node * node +| Scallmethod of expr * ident list * Int.int * expr * expr option * expr list + * node * node | Slog of expr list * expr list * node | Srevert @@ -29,7 +30,8 @@ type coq_function = { fn_return : coq_type; fn_params : (ident, coq_type) prod list; fn_temps : (ident, coq_type) prod list; fn_locals : (ident, coq_type) prod list; - fn_code : code; fn_entrypoint : node } + fn_code : code; fn_entrypoint : node; fn_nextnode : + node } val fn_return : coq_function -> coq_type @@ -43,4 +45,6 @@ val fn_code : coq_function -> code val fn_entrypoint : coq_function -> node +val fn_nextnode : coq_function -> node + type genv = (coq_function, coq_type) Genv.t diff --git a/src/backend/extraction/StmtCintptr.ml b/src/backend/extraction/StmtCintptr.ml index 900387e..4ae5818 100644 --- a/src/backend/extraction/StmtCintptr.ml +++ b/src/backend/extraction/StmtCintptr.ml @@ -18,7 +18,7 @@ type statement = | Sreturn of ident option | Shash of expr * expr * expr option | Stransfer of expr * expr -| Scallmethod of expr * ident list * Int.int * expr * expr list +| Scallmethod of expr * ident list * Int.int * expr * expr option * expr list | Slog of expr list * expr list | Srevert diff --git a/src/backend/extraction/StmtCintptr.mli b/src/backend/extraction/StmtCintptr.mli index 6f270bb..18a6812 100644 --- a/src/backend/extraction/StmtCintptr.mli +++ b/src/backend/extraction/StmtCintptr.mli @@ -18,7 +18,7 @@ type statement = | Sreturn of ident option | Shash of expr * expr * expr option | Stransfer of expr * expr -| Scallmethod of expr * ident list * Int.int * expr * expr list +| Scallmethod of expr * ident list * Int.int * expr * expr option * expr list | Slog of expr list * expr list | Srevert diff --git a/src/backend/extraction/StmtClinear.ml b/src/backend/extraction/StmtClinear.ml index 36abef7..cd6754c 100644 --- a/src/backend/extraction/StmtClinear.ml +++ b/src/backend/extraction/StmtClinear.ml @@ -19,7 +19,8 @@ type statement = | Sjumpi of expr * label | Shash of expr * expr * expr option | Stransfer of expr * expr * label -| Scallmethod of expr * ident list * Int.int * expr * expr list * label +| Scallmethod of expr * ident list * Int.int * expr * expr option * expr list + * label | Slog of expr list * expr list | Srevert | Sfetchargs of bool diff --git a/src/backend/extraction/StmtClinear.mli b/src/backend/extraction/StmtClinear.mli index f2e446c..727ff14 100644 --- a/src/backend/extraction/StmtClinear.mli +++ b/src/backend/extraction/StmtClinear.mli @@ -19,7 +19,8 @@ type statement = | Sjumpi of expr * label | Shash of expr * expr * expr option | Stransfer of expr * expr * label -| Scallmethod of expr * ident list * Int.int * expr * expr list * label +| Scallmethod of expr * ident list * Int.int * expr * expr option * expr list + * label | Slog of expr list * expr list | Srevert | Sfetchargs of bool diff --git a/src/backend/extraction/StmtClocal.ml b/src/backend/extraction/StmtClocal.ml index 6daef09..95b4bba 100644 --- a/src/backend/extraction/StmtClocal.ml +++ b/src/backend/extraction/StmtClocal.ml @@ -18,7 +18,7 @@ type statement = | Sreturn of ident option | Shash of expr * expr * expr option | Stransfer of expr * expr -| Scallmethod of expr * ident list * Int.int * expr * expr list +| Scallmethod of expr * ident list * Int.int * expr * expr option * expr list | Slog of expr list * expr list | Srevert diff --git a/src/backend/extraction/StmtClocal.mli b/src/backend/extraction/StmtClocal.mli index 62f73a3..739a2b3 100644 --- a/src/backend/extraction/StmtClocal.mli +++ b/src/backend/extraction/StmtClocal.mli @@ -18,7 +18,7 @@ type statement = | Sreturn of ident option | Shash of expr * expr * expr option | Stransfer of expr * expr -| Scallmethod of expr * ident list * Int.int * expr * expr list +| Scallmethod of expr * ident list * Int.int * expr * expr option * expr list | Slog of expr list * expr list | Srevert diff --git a/src/backend/extraction/StmtExpressionless.ml b/src/backend/extraction/StmtExpressionless.ml index a13df3c..4583979 100644 --- a/src/backend/extraction/StmtExpressionless.ml +++ b/src/backend/extraction/StmtExpressionless.ml @@ -37,7 +37,8 @@ type statement = | Sjumpi | Shash | Stransfer -| Scallmethod of Int.int * nat * nat +| Scallargs of Int.int * nat * nat +| Scallmethod of bool | Slog of nat * nat | Srevert | Scalldataload diff --git a/src/backend/extraction/StmtExpressionless.mli b/src/backend/extraction/StmtExpressionless.mli index 63a1f36..92cbe12 100644 --- a/src/backend/extraction/StmtExpressionless.mli +++ b/src/backend/extraction/StmtExpressionless.mli @@ -37,7 +37,8 @@ type statement = | Sjumpi | Shash | Stransfer -| Scallmethod of Int.int * nat * nat +| Scallargs of Int.int * nat * nat +| Scallmethod of bool | Slog of nat * nat | Srevert | Scalldataload diff --git a/src/backend/extraction/StmtMiniC.ml b/src/backend/extraction/StmtMiniC.ml index 7bdac4b..94efc28 100644 --- a/src/backend/extraction/StmtMiniC.ml +++ b/src/backend/extraction/StmtMiniC.ml @@ -16,7 +16,7 @@ type statement = | Sbreak | Sreturn of ident option | Stransfer of expr * expr -| Scallmethod of expr * ident list * Int.int * expr * expr list +| Scallmethod of expr * ident list * Int.int * expr * expr option * expr list | Slog of expr list * expr list | Srevert diff --git a/src/backend/extraction/StmtMiniC.mli b/src/backend/extraction/StmtMiniC.mli index a0208d3..8780513 100644 --- a/src/backend/extraction/StmtMiniC.mli +++ b/src/backend/extraction/StmtMiniC.mli @@ -16,7 +16,7 @@ type statement = | Sbreak | Sreturn of ident option | Stransfer of expr * expr -| Scallmethod of expr * ident list * Int.int * expr * expr list +| Scallmethod of expr * ident list * Int.int * expr * expr option * expr list | Slog of expr list * expr list | Srevert diff --git a/src/backend/extraction/StmtStacked.ml b/src/backend/extraction/StmtStacked.ml index 1882026..2beaf67 100644 --- a/src/backend/extraction/StmtStacked.ml +++ b/src/backend/extraction/StmtStacked.ml @@ -22,7 +22,8 @@ type statement = | Sjumpi | Shash | Stransfer -| Scallmethod of Int.int * nat * nat +| Scallargs of Int.int * nat * nat +| Scallmethod of bool | Slog of nat * nat | Srevert | Sfetchargs of nat diff --git a/src/backend/extraction/StmtStacked.mli b/src/backend/extraction/StmtStacked.mli index 0b5d13a..c86da1f 100644 --- a/src/backend/extraction/StmtStacked.mli +++ b/src/backend/extraction/StmtStacked.mli @@ -22,7 +22,8 @@ type statement = | Sjumpi | Shash | Stransfer -| Scallmethod of Int.int * nat * nat +| Scallargs of Int.int * nat * nat +| Scallmethod of bool | Slog of nat * nat | Srevert | Sfetchargs of nat diff --git a/src/backend/phase/MiniC/BigstepSemantics.v b/src/backend/phase/MiniC/BigstepSemantics.v new file mode 100755 index 0000000..ff4236a --- /dev/null +++ b/src/backend/phase/MiniC/BigstepSemantics.v @@ -0,0 +1,510 @@ +Require Import backend.TempModel. +Require Import backend.AST. +Require Import cclib.Coqlib. +Require Import backend.Values.HighValues. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.Cop. +Require Import backend.Ctypes. +Require Import backend.Options. +Require Import backend.MemoryModel. +Require Import backend.MachineModel. +Require Import backend.Environments.Globalenvs. +Require Import backend.Environments.AllEnvironments. +Require Import backend.Environments.ExtEnv. +Require Import backend.AbstractData. +Require Import backend.Events. +Require Import backend.GasModel. +Require Import backend.Statements.StmtMiniC. +Require Import backend.Expressions.SemanticsMiniC. +Require Import backend.phase.MiniC.Semantics. + +Section WITH_DATA. + Context (D : compatdata). + +Section BIGSTEP. + +Variable ge: genv. +Variable me: machine_env (cdata_type D). + +(** ** Big-step semantics for terminating statements and functions *) + +Inductive outcome: Type := +| Out_break: outcome (* terminated by [break] *) +| Out_normal: outcome (* terminated normally *) +| Out_return: val -> outcome. (* terminated by [return] *) + +Inductive out_normal : outcome -> Prop := +| Out_normal_N: out_normal Out_normal. + +Inductive out_break_or_return : outcome -> outcome -> Prop := +| Out_break_or_return_B: out_break_or_return Out_break Out_normal +| Out_break_or_return_R: forall ov, + out_break_or_return (Out_return ov) (Out_return ov). + +Definition outcome_switch (out: outcome) : outcome := + match out with + | Out_break => Out_normal + | o => o + end. + +(* the second log and the nat is the *new* events, and the *additional* gas used. *) +(* Inductive exec_stmt: int256 -> temp_env -> ext_env -> log -> statement -> int256 -> temp_env -> ext_env -> log -> nat -> outcome -> Prop := *) +(* | exec_Sskip: forall ctx le m lg, *) +(* exec_stmt ctx le m lg Sskip *) +(* ctx le m nil 0 Out_normal *) +(* | exec_Sassign: forall ctx le se lg lv lv_ident rv rv_int, *) +(* eval_lvalue ctx me se le lv lv_ident -> *) +(* eval_rvalue ctx me se le rv rv_int -> *) +(* True *) +(* (* exec_stmt ctx le se lg (Sassign lv rv) *) *) +(* (* ctx le (IdentExtMap.set lv_ident rv_int se) nil (gas_assign true lv rv 2) Out_normal *) *) +(* | exec_Sset: forall id rv rv_int le se prev lg, *) +(* PTree.get id le = Some prev -> *) +(* eval_rvalue me se le rv rv_int -> *) +(* True *) +(* (* exec_stmt le se lg (Sset id rv) *) *) +(* (* (PTree.set id rv_int le) se nil (gas_set true rv 2) Out_normal *) *) +(* | exec_Scall: forall optid lab fd args arg_ints vres le se se' lg lg' g, *) +(* eval_rvalues me se le args arg_ints -> *) +(* lookup_function ge lab = Some fd -> *) +(* eval_funcall se lg fd arg_ints se' lg' g vres -> *) +(* exec_stmt le se lg (Scall optid lab args) (optset optid vres le) se' lg' (g + gas_saveretval' 1 + gas_call true args 1) Out_normal *) +(* *) +(* *) +(* | exec_Sseq_1: forall le le1 le2 se se1 se2 s1 s2 lg lg1 lg2 g1 g2 out, *) +(* exec_stmt le se lg s1 le1 se1 lg1 g1 Out_normal-> *) +(* exec_stmt le1 se1 (lg1++lg) s2 le2 se2 lg2 g2 out -> *) +(* exec_stmt le se lg (Ssequence s1 s2) le2 se2 (lg2++lg1) (g1+g2) out *) +(* *) +(* | exec_Sseq_2: forall le le1 se se1 s1 s2 lg lg1 g1 out, *) +(* exec_stmt le se lg s1 le1 se1 lg1 g1 out -> *) +(* out <> Out_normal -> *) +(* exec_stmt le se lg (Ssequence s1 s2) le1 se1 lg1 g1 out *) +(* *) +(* | exec_Sifthenelse: forall le le1 se se1 cond b s1 s2 lg lg1 g1 out, *) +(* eval_rvalue me se le cond (Vint b) -> *) +(* exec_stmt le se lg (if (Int256.eq b Int256.zero) then s2 else s1) *) +(* le1 se1 lg1 g1 out -> *) +(* exec_stmt le se lg (Sifthenelse cond s1 s2) *) +(* le1 se1 lg1 (g1+gas_eval_cond true cond 3) out *) +(* *) +(* | exec_Sreturn_some: forall se le lg a v, *) +(* eval_rvalue me se le a v -> *) +(* exec_stmt le se lg (Sreturn (Some a)) *) +(* le se nil (gas_return true (Some a) 0) (Out_return v) *) +(* *) +(* | exec_Sbreak: forall se le lg, *) +(* exec_stmt le se lg Sbreak *) +(* le se nil 0 Out_break *) +(* *) +(* | exec_Sloop_stop: forall le le' se se' lg lg' s g out out', *) +(* exec_stmt le se lg s le' se' lg' g out' -> *) +(* out_break_or_return out' out -> *) +(* exec_stmt le se lg (Sloop s) *) +(* le' se' lg' (g+ gas_branches 2) out *) +(* *) +(* | exec_Sloop_loop: forall le le1 le2 se se1 se2 lg lg1 lg2 s g1 g2 out1 out2, *) +(* exec_stmt le se lg s le1 se1 lg1 g1 out1 -> *) +(* out_normal out1 -> *) +(* exec_stmt le1 se1 (lg1++lg) (Sloop s) le2 se2 lg2 g2 out2 -> *) +(* exec_stmt le se lg (Sloop s) *) +(* le2 se2 (lg2++lg1) (g1 + g2 + gas_branches 2) out2 *) +(* *) +(* | exec_Slog : forall le se topics topics' args args' lg, *) +(* eval_rvalues me se le topics topics' -> *) +(* eval_rvalues me se le args args' -> *) +(* exec_stmt le se lg (Slog topics args) le se (Elog topics' args' :: nil) (gas_log true topics args 2) Out_normal *) +(* *) +(* *) +(* (* todo: we should also have a case for Stranfer, but right now *) +(* the small-step semantics are not quite right (the relation *) +(* should also act on the storage environment), and in any *) +(* case we are not emitting it in Edsger yet, so leave it for later. *) *) +(* *) +(* (** [eval_funcall m1 fd args t m2 res] describes the invocation of *) +(* function [fd] with arguments [args]. [res] is the value returned *) +(* by the call. *) *) +(* *) +(* with eval_funcall: ext_env -> log -> function -> list val -> ext_env -> log -> nat -> val -> Prop := *) +(* | eval_funcall_internal: forall le le' se se' lg lg' f args_int g vres, *) +(* function_initial_temps ftype f args_int = Some le -> *) +(* exec_stmt le se lg f.(fn_body) le' se' lg' g (Out_return vres) -> *) +(* eval_funcall se lg f args_int se' lg' (g + gas_done (length (all_temps ftype f)) 3 + gas_callenter (length (some_temps ftype f)) (length (some_args ftype f)) 1) vres *) +(* (* *) +(* | eval_funcall_external: ... *). *) +(* *) +(* *) +(* (* In Clight, there is a distinction between internal and external function *) +(* calls, and the DeepSEA C backend generally uses external calls. *) +(* But it doesn't make a big difference, because the primitive-proofs make use of *) +(* eval_funcall, rather than eval_funcall_internal or eval_funcall_external directly. *) *) +(* *) +(* Scheme exec_stmt_ind2 := Minimality for exec_stmt Sort Prop *) +(* with eval_funcall_ind2 := Minimality for eval_funcall Sort Prop. *) +(* Combined Scheme exec_stmt_funcall_ind from exec_stmt_ind2, eval_funcall_ind2. *) + + +End BIGSTEP. + +(** Big-step execution of a whole program. *) +(* +Section WHOLE_PROGRAM. +Local Existing Instance writable_block_always_ops. + +Inductive bigstep_program_terminates' (p: program): trace -> int -> Prop := + | bigstep_program_terminates_intro: forall b f m0 m1 t r, + let ge := Genv.globalenv p in + Genv.init_mem p = Some m0 -> + Genv.find_symbol ge p.(prog_main) = Some b -> + Genv.find_funct_ptr ge b = Some f -> + type_of_fundef f = Tfunction Tnil type_int32s cc_default -> + eval_funcall ge m0 f nil t m1 (Vint r) -> + bigstep_program_terminates' p t r. + +Inductive bigstep_program_diverges' (p: program): traceinf -> Prop := + | bigstep_program_diverges_intro: forall b f m0 t, + let ge := Genv.globalenv p in + Genv.init_mem p = Some m0 -> + Genv.find_symbol ge p.(prog_main) = Some b -> + Genv.find_funct_ptr ge b = Some f -> + type_of_fundef f = Tfunction Tnil type_int32s cc_default -> + evalinf_funcall ge m0 f nil t -> + bigstep_program_diverges' p t. + +Definition bigstep_semantics' (p: program) := + Bigstep_semantics (bigstep_program_terminates' p) (bigstep_program_diverges' p). + +Definition bigstep_semantics1 := bigstep_semantics' function_entry1. + +Definition bigstep_semantics2 := bigstep_semantics' (fun _ => function_entry2). + +End WHOLE_PROGRAM. +*) + +(** * Implication from big-step semantics to transition semantics *) + +Section BIGSTEP_TO_TRANSITIONS. + +Variable ge: genv. +Variable me: machine_env (cdata_type D). + +Open Scope nat_scope. + +(* Inductive outcome_state_match *) +(* (le: temp_env) (d: cdata_type D) (se : storage_env) (f: function) (k: cont) (g:nat): *) +(* outcome -> state -> Prop := *) +(* | osm_normal: forall g', *) +(* g = g' -> *) +(* outcome_state_match le d se f k g Out_normal (State f Sskip k le d se g') *) +(* | osm_break: forall g', *) +(* g = g' -> *) +(* outcome_state_match le d se f k g Out_break (State f Sbreak k le d se g') *) +(* (* | osm_return_none: forall g' k', *) +(* lg = lg' -> *) +(* g = g' -> *) +(* call_cont k' = call_cont k -> *) +(* outcome_state_match le m f k g *) +(* (Out_return None) (State f (Sreturn None) k' le m g') *) *) +(* | osm_return_some: forall a v g' k', *) +(* g + (gas_return true (Some a) 0) = g' -> *) +(* call_cont k' = call_cont k -> *) +(* eval_rvalue me se le a v -> *) +(* outcome_state_match le d se f k g *) +(* (Out_return v) *) +(* (State f (Sreturn (Some a)) k' le d se g'). *) +(* *) +Lemma is_call_cont_call_cont: + forall k, is_call_cont k -> call_cont k = k. +Proof. + destruct k; simpl; intros; contradiction || auto. +Qed. + +(*Notation step := (fun ge => Clight.step ge (function_entry ge)).*) + +Require Import Smallstep. + + +Lemma gt_plus : forall G g, + (G >= g -> exists g0, G = g0 + g )%nat. +Proof. + intros. + exists (G-g). + omega. +Qed. + +Ltac to_plus := + match goal with + [ H : (_ >= _)%nat |- _] => (destruct (gt_plus _ _ H); subst) + end. + + +Lemma loop_inversion1 : forall f s k le d se g S2, + step me ge + (State f (Sloop s) k le d se g) S2 -> + exists g', + star (step me) ge + (State f s (Kloop s k) le d se g') S2 + /\ g' = g - gas_branches 2 + /\ g = g-gas_branches 2+gas_branches 2. +Proof. + intros. + inversion H. + subst. + eexists. + split. + - apply star_refl. + - simpl. omega. +Qed. + +(* Lemma loop_inversion : forall f s k le se lg g S2 *) +(* f' k' le' se' lg' g' out', *) +(* star (step me) ge (State f (Sloop s) k le se lg g) S2 -> *) +(* outcome_state_match le' se' f' k' lg' g' out' S2 -> *) +(* exists g', *) +(* star (step me) ge (State f s (Kloop s k) le se lg g') S2 *) +(* /\ g' = g-gas_branches 2 *) +(* /\ g = g-gas_branches 2+gas_branches 2. *) +(* Proof. *) +(* intros f s k le se lg g S2 *) +(* f' k' le' se' lg' g' out' Hstar Hosm. *) +(* inv Hstar. *) +(* + inv Hosm. *) +(* + destruct (loop_inversion1 _ _ _ _ _ _ _ _ H) as [g1 [Hstep' [Heq_g1 Heq_g]]]. *) +(* exists g1. *) +(* split; auto. *) +(* apply star_trans with s2; assumption. *) +(* Qed. *) + +(* Lemma exec_stmt_eval_funcall_steps: *) +(* (forall le m s le' m' g out, *) +(* exec_stmt ge me le m s le' m' g out -> *) +(* forall f k G, *) +(* (G >= g) -> *) +(* exists S, *) +(* star (step me) ge (State f s k le (snd m) (fst m) G) S *) +(* /\ outcome_state_match le' (snd m') (fst m') f k (G-g) out S) *) +(* /\ *) +(* (forall m f args m' g res, *) +(* eval_funcall ge me m f args m' g res -> *) +(* forall k G, *) +(* (G >= g) -> *) +(* is_call_cont k -> *) +(* star (step me) ge (Callstate f args k (snd m) (fst m) G) *) +(* (Returnstate res k (snd m') (fst m') (G - g))). *) +(* Proof. *) +(* apply exec_stmt_funcall_ind; intros. *) +(* *) +(* (* skip *) *) +(* - econstructor; split. apply star_refl. *) +(* replace (G-0)%nat with G by omega. *) +(* constructor; simpl; auto. *) +(* *) +(* (* assign *) *) +(* - intros. *) +(* to_plus. *) +(* econstructor; split. apply star_one. *) +(* econstructor; eauto. *) +(* constructor; auto; omega. *) +(* (* set *) *) +(* - intros. *) +(* to_plus. *) +(* econstructor; split. apply star_one. *) +(* econstructor; eauto. constructor; auto; omega. *) +(* *) +(* (* call *) *) +(* - intros. *) +(* to_plus. *) +(* econstructor; split. *) +(* eapply star_step. *) +(* rewrite !plus_assoc. *) +(* econstructor; eauto. *) +(* eapply star_right. *) +(* + eapply H2. *) +(* * omega. *) +(* * exact I. *) +(* + replace (x + g + gas_saveretval' 1 - g) with (x + gas_saveretval' 1) by omega. *) +(* apply step_returnstate. *) +(* + constructor; auto; omega. *) +(* *) +(* (* sequence 1 *) *) +(* - to_plus. rewrite !plus_assoc. *) +(* destruct (H0 f (Kseq s2 k) (x+g1+g2)) as [S1 [A1 B1]]. *) +(* { omega. } *) +(* inv B1. *) +(* destruct (H2 f k (x+g2)) as [S2 [A2 B2]]. *) +(* { omega. } *) +(* econstructor; split. *) +(* eapply star_left. econstructor. *) +(* eapply star_trans. eexact A1. *) +(* eapply star_left. constructor. *) +(* replace (x + g1 + g2 - g1) with (x+g2) by omega. eexact A2. *) +(* replace (x + g1 + g2 - (g1 + g2)) with (x + g2 - g2) by omega. *) +(* exact B2. *) +(* *) +(* (* sequence 2 *) *) +(* - to_plus. *) +(* destruct (H0 f (Kseq s2 k) (x+g1)) as [S1 [A1 B1]]. *) +(* { omega. } *) +(* set (S2 := *) +(* match out with *) +(* | Out_break => State f Sbreak k le1 (snd m1) (fst m1) x *) +(* | _ => S1 *) +(* end). *) +(* exists S2; split. *) +(* + eapply star_left. econstructor. *) +(* eapply star_trans. eexact A1. *) +(* unfold S2; inv B1. *) +(* congruence. *) +(* * apply star_one. *) +(* replace (x + g1 - g1) with x by omega. *) +(* apply step_break_seq. *) +(* * apply star_refl. *) +(* + unfold S2; inv B1; congruence || econstructor; eauto; omega. *) +(* *) +(* (* ifthenelse *) *) +(* - to_plus. *) +(* destruct (H1 f k (x+g1)) as [S1 [A1 B1]]. *) +(* { omega. } *) +(* exists S1; split. *) +(* + eapply star_left. 2: eexact A1. *) +(* rewrite plus_assoc. *) +(* eapply step_ifthenelse; eauto. *) +(* + replace(x + (g1 + gas_eval_cond true cond 3) - (g1 + gas_eval_cond true cond 3)) *) +(* with (x + g1 - g1) by omega. *) +(* exact B1. *) +(* *) +(* (* return none *) *) +(* (* *) +(* - econstructor; split. apply star_refl. *) +(* apply osm_return_none; simpl; auto; omega. *) +(* *) *) +(* *) +(* *) +(* (* return some *) *) +(* -econstructor; split. apply star_refl. *) +(* apply osm_return_some; simpl; auto; omega. *) +(* *) +(* *) +(* (* break *) *) +(* - econstructor; split. apply star_refl. *) +(* apply osm_break; simpl; auto; omega. *) +(* *) +(* (* loop stop *) *) +(* - to_plus. rewrite !plus_assoc. *) +(* destruct (H0 f (Kloop s k) (x+g)) as [S1 [A1 B1]]. *) +(* { omega. } *) +(* set (S2 := *) +(* match out' with *) +(* | Out_break => State f Sskip k le' (snd m') (fst m') (x) *) +(* | _ => S1 *) +(* end). *) +(* exists S2; split. *) +(* + eapply star_left. eapply step_loop. *) +(* eapply star_trans. eexact A1. *) +(* unfold S2. inversion H1; subst. *) +(* inv B1. apply star_one. *) +(* replace (x+g-g) with x by omega. *) +(* constructor. *) +(* apply star_refl. *) +(* + replace (x + g + gas_branches 2 - (g + gas_branches 2)) *) +(* with (x + g - g) by omega. *) +(* unfold S2. inversion H1; subst. *) +(* * constructor; auto; omega. *) +(* * inversion B1; apply osm_return_some; simpl in *; auto. *) +(* *) +(* (* loop loop *) *) +(* - to_plus. *) +(* rewrite !plus_assoc. *) +(* destruct (H0 f (Kloop s k) (x+g1+g2)) as [S1 [A1 B1]]. *) +(* { omega. } *) +(* destruct (H3 f k (x+g2)) as [S2 [A2 B2]]. *) +(* { omega. } *) +(* exists S2; split. *) +(* + eapply star_left. eapply step_loop. *) +(* eapply star_trans. *) +(* exact A1. *) +(* *) +(* eapply loop_inversion in A2; [|eassumption]. *) +(* destruct A2 as [g' [A2_steps [A2_eq1 A2_eq2]]]. *) +(* inv H1. *) +(* inv B1. *) +(* eapply star_left. *) +(* replace (x + g1 + g2 - g1) with (x + g2) by omega. *) +(* rewrite A2_eq2. *) +(* apply step_skip_loop. *) +(* exact A2_steps. *) +(* + replace (x + g1 + g2 + gas_branches 2 - (g1 + g2 + gas_branches 2)) *) +(* with (x + g2 - g2) by omega. *) +(* inv B2; constructor; auto. *) +(* *) +(* (* log *) *) +(* - to_plus. *) +(* eexists. split. *) +(* *) +(* +eapply star_step. *) +(* econstructor. *) +(* eauto. *) +(* eauto. *) +(* econstructor. *) +(* + constructor; simpl; auto. *) +(* omega. *) +(* *) +(* (* call internal *) *) +(* - to_plus. *) +(* rewrite !plus_assoc. *) +(* destruct (H1 f k (x + g + gas_done (length (all_temps ftype f)) 3)) as [S1 [A1 B1]]. *) +(* { omega. } *) +(* eapply star_left. eapply step_internal_function; eauto. *) +(* eapply star_right. eexact A1. *) +(* inv B1. *) +(* (* Out_return Some *) *) +(* + rewrite <- (is_call_cont_call_cont k H3). rewrite <- H6. *) +(* replace *) +(* (x + g + gas_done (length (all_temps ftype f)) 3 + *) +(* gas_callenter (length (some_temps ftype f)) (length (some_args ftype f)) 1 - *) +(* (g + gas_done (length (all_temps ftype f)) 3 + *) +(* gas_callenter (length (some_temps ftype f)) (length (some_args ftype f)) 1)) *) +(* with x by omega. *) +(* replace (x + g + gas_done (length (all_temps ftype f)) 3 - g + gas_return true (Some a) 0) *) +(* with (x + (gas_done (length (all_temps ftype f)) 3 + gas_return true (Some a) 0)) *) +(* by omega. *) +(* replace (gas_done (length (all_temps ftype f)) 3 + gas_return true (Some a) 0) *) +(* with (gas_return_done true (Some a) (length (all_temps ftype f)) 3). *) +(* *) +(* Focus 2. *) +(* rewrite <- gas_return_done_split; f_equal; omega. *) +(* eapply step_return; eauto. *) +(* constructor; assumption. *) +(* (* call external *) *) +(* (* - apply star_one. apply step_external_function; auto. *) *) +(* Qed. *) + + +(* Lemma exec_stmt_steps: *) +(* forall le m s le' m' g out, *) +(* exec_stmt ge me le m s le' m' g out -> *) +(* forall f k G, *) +(* (G >= g) -> *) +(* exists S, *) +(* star (step me) ge (State f s k le (snd m) (fst m) G) S *) +(* /\ outcome_state_match le' (snd m') (fst m') f k (G-g) out S. *) +(* Proof. *) +(* destruct exec_stmt_eval_funcall_steps; assumption. *) +(* Qed. *) +(* *) +(* Lemma eval_funcall_steps: *) +(* forall m f args m' g res, *) +(* eval_funcall ge me m f args m' g res -> *) +(* forall k G, *) +(* (G >= g) -> *) +(* is_call_cont k -> *) +(* star (step me) ge (Callstate f args k (snd m) (fst m) G) *) +(* (Returnstate res k (snd m') (fst m') (G-g)). *) +(* Proof. *) +(* destruct exec_stmt_eval_funcall_steps; assumption. *) +(* Qed. *) + +End BIGSTEP_TO_TRANSITIONS. + +End WITH_DATA. diff --git a/src/backend/phase/MiniC/Semantics.v b/src/backend/phase/MiniC/Semantics.v new file mode 100755 index 0000000..b64d09f --- /dev/null +++ b/src/backend/phase/MiniC/Semantics.v @@ -0,0 +1,264 @@ +(* Semantics for MiniC expressions and statements. *) + +Require Import backend.TempModel. +Require Import backend.AST. +Require Import cclib.Coqlib. +Require Import backend.Expressions.ExpMiniC. +Require Import backend.Expressions.SemanticsMiniC. +Require Import backend.Statements.StmtMiniC. +Require Import backend.Values.HighValues. +Require Import cclib.Maps. +Require Import cclib.Integers. +Require Import backend.Cop. +Require Import backend.Ctypes. +Require Import backend.Options. +Require Import backend.MemoryModel. +Require Import backend.AbstractData. +Require Import backend.MachineModel. +Require Import backend.Environments.Globalenvs. +Require Import backend.Environments.AllEnvironments. +Require Import backend.Environments.ExtEnv. +Require Import backend.GasModel. +Require Import backend.SymbolicKeccak. +Require Import List. Import ListNotations. + +Definition ftype : ftemps := mkftemps function fn_temps fn_params. + +Section WITH_DATA. + + +(* machine environment *) +Context {adata} {data_ops: CompatDataOps adata}. +Variable me: machine_env adata. + +(** ** Transition semantics for statements and functions *) + +(** Continuations *) + +(* from a given backtrace / line number in the program, where does it go from here? +forms a linked list of continuations, ending at a stop. *) + +Inductive cont: Type := +| Kstop: cont +| Kseq: statement -> cont -> cont (**r [Kseq s2 k] = after [s1] in [s1;s2] *) +| Kloop: statement -> cont -> cont (**r [Kloop s k] = inside [Sloop s] *) +| Kcall: option ident -> (**r where to store result *) + function -> (**r calling function *) + temp_env -> (**r temporary env of calling function *) + cont -> cont. + +(** Pop continuation until a call or stop. + It used in the semantics on return, to pop out of all sequences and loops. + *) +Fixpoint call_cont (k: cont) : cont := + match k with + | Kseq _ k => call_cont k + | Kloop _ k => call_cont k + | _ => k + end. + +(* Used in theorem statements. *) +Definition is_call_cont (k: cont) : Prop := + match k with + | Kstop => True + | Kcall _ _ _ _ => True + | _ => False + end. + +Lemma call_cont_idempotent: forall k, + is_call_cont k -> k = call_cont k. +Proof. +intros. destruct k; simpl. +reflexivity. +simpl in H. contradiction. +simpl in H. contradiction. +reflexivity. +Qed. + +Fixpoint call_cont_index (k:cont) : int256 := + match k with + | Kstop => Int256.repr 0 + | Kcall _ _ _ k => Int256.add (call_cont_index k) (Int256.repr 1) + | Kseq _ k => call_cont_index k + | Kloop _ k => call_cont_index k + end. + +(** States *) + +Inductive state: Type := + | State + (f: function) + (s: statement) + (k: cont) + (le: temp_env) + (ee: ext_env) (* EVM specific *) + (d: adata) + (gas: nat) : state (* EVM specific. This is the amount of gas _remaining_. *) + | Callstate (* these are used for internal function calls, but currently not for method calls. *) + (fd: function) + (args: list val) + (k: cont) + (ee: ext_env) + (d: adata) + (gas: nat) : state + | Initialstate + (fd: function) + (args: list val) + (ee: ext_env) + (d: adata) + (gas: nat) : state + | Returnstate + (res: val) + (k: cont) + (ee: ext_env) + (d: adata) + (gas: nat) : state. + +Section STEP. + + Local Open Scope nat_scope. + + Variable ge: genv. + + Definition count_methods := Genv.count_methods function type ge. + + Definition lookup_function (lab: label) : option function := + PTree.get lab (Genv.genv_fundefs function type ge). + + (* This models sucessful runs. In particular, note that there is no rule for revert, so a call for revert is a stuck state. *) + Inductive step: state -> state -> Prop := + | step_assign: forall f e1 lv e2 rv k ee le d g ee', + eval_lvalue (call_cont_index k) me ee le e1 lv -> + eval_rvalue (call_cont_index k) me ee le e2 rv -> + (* only assign to locations that are already defined *) + write lv rv ee = Some ee' -> + step (State f (Sassign e1 e2) k le ee d (g + gas_assign true e1 e2 2)) + (* See GasModel.v for the definitions of gas costs. *) + (State f Sskip k le ee' d g) + + | step_set: forall f id rv rv_int k le ee prev d g, + PTree.get id le = Some prev -> (* Can only Sset an existing temp *) + eval_rvalue (call_cont_index k) me ee le rv rv_int -> + step (State f (Sset id rv) k le ee d (g + gas_set true rv 2)) + (State f Sskip k (PTree.set id rv_int le) ee d g) + + | step_call: forall f optid lab fd args arg_ints k le ee d g, + eval_rvalues (call_cont_index k) me ee le args arg_ints -> + lookup_function lab = Some fd -> + (* gas for returning is reserved in step_returnstate *) + step (State f (Scall optid lab args) k le ee d (g + gas_call true args 1)) + (Callstate fd arg_ints (Kcall optid f le k) ee d g) + + | step_seq: forall f s1 s2 k le ee d g, + step (State f (Ssequence s1 s2) k le ee d g) + (State f s1 (Kseq s2 k) le ee d g) + | step_skip_seq: forall f s k le ee d g, + step (State f Sskip (Kseq s k) le ee d g) + (State f s k le ee d g) + | step_break_seq: forall f s k le ee d g, + step (State f Sbreak (Kseq s k) le ee d g) + (State f Sbreak k le ee d g) + + | step_ifthenelse: forall f cond b s1 s2 k le ee d g, + eval_rvalue (call_cont_index k) me ee le cond (Vint b) -> + step (State f (Sifthenelse cond s1 s2) k le ee d (g + gas_eval_cond true cond 3)) + (State f (if (Int256.eq b Int256.zero) then s2 else s1) k le ee d g) + + | step_loop: forall f s k le ee d g, + (* enter loop *) + step (State f (Sloop s) k le ee d (g + gas_branches 2)) + (State f s (Kloop s k) le ee d g) + | step_break_loop: forall f s k le ee d g, + (* break from any part of the loop. doesn't use gas because it doesn't compile to any steps in Cgraph *) + step (State f Sbreak (Kloop s k) le ee d g) + (State f Sskip k le ee d g) + | step_skip_loop: forall f s k le ee d g, + (* reach the end of s, go back to restart the loop *) + step (State f Sskip (Kloop s k) le ee d (g + gas_branches 2)) + (State f s (Kloop s k) le ee d g) + + | step_return_var: forall f retvar val k le ee ee' ctx' d g, + PTree.get retvar le = Some val -> + push_func function fn_locals (call_cont_index k) f ee = Some (ee',ctx') -> + step (State f (Sreturn (Some retvar)) k le ee d (g + gas_return_done true (Some retvar) (length (all_temps ftype f)) 3)) + (* TODO: gas for pushing the retval *) + (Returnstate val (call_cont k) ee' d g) + | step_return_none: forall f k le ee ee' ctx' d g, + push_func function fn_locals (call_cont_index k) f ee = Some (ee',ctx') -> + step (State f (Sreturn None) k le ee d (g + gas_return_done true None (length (all_temps ftype f)) 3)) + (* TODO: gas for pushing the retval *) + (Returnstate Vunit (call_cont k) ee' d g) + + + | step_internal_function: forall f arg_ints k le ee ee' ctx' d g, + function_initial_temps ftype f arg_ints = Some le -> + push_func function fn_locals (call_cont_index k) f ee = Some (ee',ctx') -> + step (Callstate f arg_ints k ee d (g + gas_callenter (length (some_temps ftype f)) (length (some_args ftype f)) 1)) + (State f (fn_body f) k le ee d g) + + | step_returnstate: forall retval id f k le ee d g, + step (Returnstate retval (Kcall id f le k) ee d (g + gas_saveretval' 1)) + (State f Sskip k (optset id retval le) ee d g) + + | step_transfer: forall f a a' v v' le d d' ee k g, + eval_rvalue (call_cont_index k) me ee le a a' -> + eval_rvalue (call_cont_index k) me ee le v v' -> + (me_transfer me) a' v' d d' Int256.one -> + step (State f (Stransfer a v) k le ee d (g + gas_transfer true a v 2)) + (State f Sskip k le ee d' g) + + | step_log: forall f topics topics' args args' le d ee k g, + eval_rvalues (call_cont_index k) me ee le topics topics' -> + eval_rvalues (call_cont_index k) me ee le args args' -> + step (State f (Slog topics args) k le ee d (g + gas_log true topics args 2)) + (State f Sskip k le ee (me_log me topics' args' d) g) + + (* This currently only defines what happens if the method succeeds. + + Whether this is a problem depends on what we do in the source + language: if we don't provide a way to catch failures, then it + makes sense to only model succeeding runs (and in practice + Solidity programs may not catch failures.) + + In our compiler, we have a test for a failing method call, and if so we call revert. + *) + | step_callmethod: forall f a a' sg v v' args args' le le' ee ee' d d' k rvs rv_keys g gas, + eval_rvalue (call_cont_index k) me ee le a a' -> + eval_rvalue (call_cont_index k) me ee le v v' -> + eval_rvalues (call_cont_index k) me ee le args args' -> + (me_callmethod me) a' sg v' args' d ee d' ee' Int256.one rvs -> + listset rv_keys rvs le = Some le' -> + step (State f (Scallmethod a rv_keys sg v gas args) k le ee d (g + gas_callmethod true a (length rv_keys) v args 2)) + (State f Sskip k le' ee d' g). + + Print step. + (* Whole-program semantics + A simulation proof consists of three parts: the steps match, the initial states match, and the final states match. + Here we define the initial and final states. + *) + + (* Note that these take more arguments than Clight. *) + + Inductive initial_state (sg: int) (*our method signature, this is provided to us by the Solidity ABI.*) + (args: list val) + (d: adata) + (ee: ext_env) + (g: nat): state -> Prop := + | initial_state_intro: forall f, + In sg (Genv.genv_methods function type ge) -> + IntMap.get sg (Genv.genv_methoddefs function type ge) = Some f -> + initial_state sg args d ee g (Initialstate f args ee d g). + + (* note we may care how much gas is left, if this is called from another contract *) + Inductive final_state: state -> val -> adata -> ext_env -> nat -> Prop := + | final_state_intro: forall r ee d g, + final_state (Returnstate r Kstop d ee g) r ee d g. + +End STEP. + +End WITH_DATA. + +Arguments step {adata}. +Arguments state {adata}. +Arguments initial_state {adata}. +Arguments final_state {adata}. diff --git a/src/cclib/Coqlib.v b/src/cclib/Coqlib.v new file mode 100755 index 0000000..18d4d7e --- /dev/null +++ b/src/cclib/Coqlib.v @@ -0,0 +1,1433 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** This file collects a number of definitions and theorems that are + used throughout the development. It complements the Coq standard + library. *) + +Require Export ZArith. +Require Export Znumtheory. +Require Export List. +Require Export Bool. + +Global Set Asymmetric Patterns. + +(** * Useful tactics *) + +Ltac inv H := inversion H; clear H; subst. + +Ltac predSpec pred predspec x y := + generalize (predspec x y); case (pred x y); intro. + +Ltac caseEq name := + generalize (refl_equal name); pattern name at -1 in |- *; case name. + +Ltac destructEq name := + destruct name eqn:?. + +Ltac decEq := + match goal with + | [ |- _ = _ ] => f_equal + | [ |- (?X ?A <> ?X ?B) ] => + cut (A <> B); [intro; congruence | try discriminate] + end. + +Ltac byContradiction := + cut False; [contradiction|idtac]. + +Ltac omegaContradiction := + cut False; [contradiction|omega]. + +Lemma modusponens: forall (P Q: Prop), P -> (P -> Q) -> Q. +Proof. auto. Qed. + +Ltac exploit x := + refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _ _) _) + || refine (modusponens _ _ (x _ _ _) _) + || refine (modusponens _ _ (x _ _) _) + || refine (modusponens _ _ (x _) _). + +(** * Definitions and theorems over the type [positive] *) + +Definition peq: forall (x y: positive), {x = y} + {x <> y} := Pos.eq_dec. +Global Opaque peq. + +Lemma peq_true: + forall (A: Type) (x: positive) (a b: A), (if peq x x then a else b) = a. +Proof. + intros. case (peq x x); intros. + auto. + elim n; auto. +Qed. + +Lemma peq_false: + forall (A: Type) (x y: positive) (a b: A), x <> y -> (if peq x y then a else b) = b. +Proof. + intros. case (peq x y); intros. + elim H; auto. + auto. +Qed. + +Definition Plt: positive -> positive -> Prop := Pos.lt. + +Lemma Plt_ne: + forall (x y: positive), Plt x y -> x <> y. +Proof. + unfold Plt; intros. red; intro. subst y. eelim Pos.lt_irrefl; eauto. +Qed. +Hint Resolve Plt_ne: coqlib. + +Lemma Plt_trans: + forall (x y z: positive), Plt x y -> Plt y z -> Plt x z. +Proof (Pos.lt_trans). + +Lemma Plt_succ: + forall (x: positive), Plt x (Psucc x). +Proof. + unfold Plt; intros. apply Pos.lt_succ_r. apply Pos.le_refl. +Qed. +Hint Resolve Plt_succ: coqlib. + +Lemma Plt_trans_succ: + forall (x y: positive), Plt x y -> Plt x (Psucc y). +Proof. + intros. apply Plt_trans with y. assumption. apply Plt_succ. +Qed. +Hint Resolve Plt_succ: coqlib. + +Lemma Plt_succ_inv: + forall (x y: positive), Plt x (Psucc y) -> Plt x y \/ x = y. +Proof. + unfold Plt; intros. rewrite Pos.lt_succ_r in H. + apply Pos.le_lteq; auto. +Qed. + +Definition plt (x y: positive) : {Plt x y} + {~ Plt x y}. +Proof. + unfold Plt, Pos.lt; intros. destruct (Pos.compare x y). + - right; congruence. + - left; auto. + - right; congruence. +Defined. +Global Opaque plt. + +Definition Ple: positive -> positive -> Prop := Pos.le. + +Lemma Ple_refl: forall (p: positive), Ple p p. +Proof (Pos.le_refl). + +Lemma Ple_trans: forall (p q r: positive), Ple p q -> Ple q r -> Ple p r. +Proof (Pos.le_trans). + +Lemma Plt_Ple: forall (p q: positive), Plt p q -> Ple p q. +Proof (Pos.lt_le_incl). + +Lemma Ple_succ: forall (p: positive), Ple p (Psucc p). +Proof. + intros. apply Plt_Ple. apply Plt_succ. +Qed. + +Lemma Plt_Ple_trans: + forall (p q r: positive), Plt p q -> Ple q r -> Plt p r. +Proof (Pos.lt_le_trans). + +Lemma Plt_strict: forall p, ~ Plt p p. +Proof (Pos.lt_irrefl). + +Hint Resolve Ple_refl Plt_Ple Ple_succ Plt_strict: coqlib. + +Ltac xomega := unfold Plt, Ple in *; zify; omega. +Ltac xomegaContradiction := exfalso; xomega. + +(** Peano recursion over positive numbers. *) + +Section POSITIVE_ITERATION. + +Lemma Plt_wf: well_founded Plt. +Proof. + apply well_founded_lt_compat with nat_of_P. + intros. apply nat_of_P_lt_Lt_compare_morphism. exact H. +Qed. + +Variable A: Type. +Variable v1: A. +Variable f: positive -> A -> A. + +Lemma Ppred_Plt: + forall x, x <> xH -> Plt (Ppred x) x. +Proof. + intros. elim (Psucc_pred x); intro. contradiction. + set (y := Ppred x) in *. rewrite <- H0. apply Plt_succ. +Qed. + +Let iter (x: positive) (P: forall y, Plt y x -> A) : A := + match peq x xH with + | left EQ => v1 + | right NOTEQ => f (Ppred x) (P (Ppred x) (Ppred_Plt x NOTEQ)) + end. + +Definition positive_rec : positive -> A := + Fix Plt_wf (fun _ => A) iter. + +Lemma unroll_positive_rec: + forall x, + positive_rec x = iter x (fun y _ => positive_rec y). +Proof. + unfold positive_rec. apply (Fix_eq Plt_wf (fun _ => A) iter). + intros. unfold iter. case (peq x 1); intro. auto. decEq. apply H. +Qed. + +Lemma positive_rec_base: + positive_rec 1%positive = v1. +Proof. + rewrite unroll_positive_rec. unfold iter. case (peq 1 1); intro. + auto. elim n; auto. +Qed. + +Lemma positive_rec_succ: + forall x, positive_rec (Psucc x) = f x (positive_rec x). +Proof. + intro. rewrite unroll_positive_rec. unfold iter. + case (peq (Psucc x) 1); intro. + destruct x; simpl in e; discriminate. + rewrite Ppred_succ. auto. +Qed. + +Lemma positive_Peano_ind: + forall (P: positive -> Prop), + P xH -> + (forall x, P x -> P (Psucc x)) -> + forall x, P x. +Proof. + intros. + apply (well_founded_ind Plt_wf P). + intros. + case (peq x0 xH); intro. + subst x0; auto. + elim (Psucc_pred x0); intro. contradiction. rewrite <- H2. + apply H0. apply H1. apply Ppred_Plt. auto. +Qed. + +End POSITIVE_ITERATION. + +(** * Definitions and theorems over the type [Z] *) + +Definition zeq: forall (x y: Z), {x = y} + {x <> y} := Z.eq_dec. + +Lemma zeq_true: + forall (A: Type) (x: Z) (a b: A), (if zeq x x then a else b) = a. +Proof. + intros. case (zeq x x); intros. + auto. + elim n; auto. +Qed. + +Lemma zeq_false: + forall (A: Type) (x y: Z) (a b: A), x <> y -> (if zeq x y then a else b) = b. +Proof. + intros. case (zeq x y); intros. + elim H; auto. + auto. +Qed. + +Open Scope Z_scope. + +Definition zlt: forall (x y: Z), {x < y} + {x >= y} := Z_lt_dec. + +Lemma zlt_true: + forall (A: Type) (x y: Z) (a b: A), + x < y -> (if zlt x y then a else b) = a. +Proof. + intros. case (zlt x y); intros. + auto. + omegaContradiction. +Qed. + +Lemma zlt_false: + forall (A: Type) (x y: Z) (a b: A), + x >= y -> (if zlt x y then a else b) = b. +Proof. + intros. case (zlt x y); intros. + omegaContradiction. + auto. +Qed. + +Definition zle: forall (x y: Z), {x <= y} + {x > y} := Z_le_gt_dec. + +Lemma zle_true: + forall (A: Type) (x y: Z) (a b: A), + x <= y -> (if zle x y then a else b) = a. +Proof. + intros. case (zle x y); intros. + auto. + omegaContradiction. +Qed. + +Lemma zle_false: + forall (A: Type) (x y: Z) (a b: A), + x > y -> (if zle x y then a else b) = b. +Proof. + intros. case (zle x y); intros. + omegaContradiction. + auto. +Qed. + +(** Properties of powers of two. *) + +Lemma two_power_nat_O : two_power_nat O = 1. +Proof. reflexivity. Qed. + +Lemma two_power_nat_pos : forall n : nat, two_power_nat n > 0. +Proof. + induction n. rewrite two_power_nat_O. omega. + rewrite two_power_nat_S. omega. +Qed. + +Lemma two_power_nat_two_p: + forall x, two_power_nat x = two_p (Z_of_nat x). +Proof. + induction x. auto. + rewrite two_power_nat_S. rewrite inj_S. rewrite two_p_S. omega. omega. +Qed. + +Lemma two_p_monotone: + forall x y, 0 <= x <= y -> two_p x <= two_p y. +Proof. + intros. + replace (two_p x) with (two_p x * 1) by omega. + replace y with (x + (y - x)) by omega. + rewrite two_p_is_exp; try omega. + apply Zmult_le_compat_l. + assert (two_p (y - x) > 0). apply two_p_gt_ZERO. omega. omega. + assert (two_p x > 0). apply two_p_gt_ZERO. omega. omega. +Qed. + +Lemma two_p_monotone_strict: + forall x y, 0 <= x < y -> two_p x < two_p y. +Proof. + intros. assert (two_p x <= two_p (y - 1)). apply two_p_monotone; omega. + assert (two_p (y - 1) > 0). apply two_p_gt_ZERO. omega. + replace y with (Zsucc (y - 1)) by omega. rewrite two_p_S. omega. omega. +Qed. + +Lemma two_p_strict: + forall x, x >= 0 -> x < two_p x. +Proof. + intros x0 GT. pattern x0. apply natlike_ind. + simpl. omega. + intros. rewrite two_p_S; auto. generalize (two_p_gt_ZERO x H). omega. + omega. +Qed. + +Lemma two_p_strict_2: + forall x, x >= 0 -> 2 * x - 1 < two_p x. +Proof. + intros. assert (x = 0 \/ x - 1 >= 0) by omega. destruct H0. + subst. vm_compute. auto. + replace (two_p x) with (2 * two_p (x - 1)). + generalize (two_p_strict _ H0). omega. + rewrite <- two_p_S. decEq. omega. omega. +Qed. + +(** Properties of [Zmin] and [Zmax] *) + +Lemma Zmin_spec: + forall x y, Zmin x y = if zlt x y then x else y. +Proof. + intros. case (zlt x y); unfold Zlt, Zge; intro z. + unfold Zmin. rewrite z. auto. + unfold Zmin. caseEq (x ?= y); intro. + apply Zcompare_Eq_eq. auto. + contradiction. + reflexivity. +Qed. + +Lemma Zmax_spec: + forall x y, Zmax x y = if zlt y x then x else y. +Proof. + intros. case (zlt y x); unfold Zlt, Zge; intro z. + unfold Zmax. rewrite <- (Zcompare_antisym y x). + rewrite z. simpl. auto. + unfold Zmax. rewrite <- (Zcompare_antisym y x). + caseEq (y ?= x); intro; simpl. + symmetry. apply Zcompare_Eq_eq. auto. + contradiction. reflexivity. +Qed. + +Lemma Zmax_bound_l: + forall x y z, x <= y -> x <= Zmax y z. +Proof. + intros. generalize (Zmax1 y z). omega. +Qed. +Lemma Zmax_bound_r: + forall x y z, x <= z -> x <= Zmax y z. +Proof. + intros. generalize (Zmax2 y z). omega. +Qed. + +(** Properties of Euclidean division and modulus. *) + +Lemma Zdiv_small: + forall x y, 0 <= x < y -> x / y = 0. +Proof. + intros. assert (y > 0). omega. + assert (forall a b, + 0 <= a < y -> + 0 <= y * b + a < y -> + b = 0). + intros. + assert (b = 0 \/ b > 0 \/ (-b) > 0). omega. + elim H3; intro. + auto. + elim H4; intro. + assert (y * b >= y * 1). apply Zmult_ge_compat_l. omega. omega. + omegaContradiction. + assert (y * (-b) >= y * 1). apply Zmult_ge_compat_l. omega. omega. + rewrite <- Zopp_mult_distr_r in H6. omegaContradiction. + apply H1 with (x mod y). + apply Z_mod_lt. auto. + rewrite <- Z_div_mod_eq. auto. auto. +Qed. + +Lemma Zmod_small: + forall x y, 0 <= x < y -> x mod y = x. +Proof. + intros. assert (y > 0). omega. + generalize (Z_div_mod_eq x y H0). + rewrite (Zdiv_small x y H). omega. +Qed. + +Lemma Zmod_unique: + forall x y a b, + x = a * y + b -> 0 <= b < y -> x mod y = b. +Proof. + intros. subst x. rewrite Zplus_comm. + rewrite Z_mod_plus. apply Zmod_small. auto. omega. +Qed. + +Lemma Zdiv_unique: + forall x y a b, + x = a * y + b -> 0 <= b < y -> x / y = a. +Proof. + intros. subst x. rewrite Zplus_comm. + rewrite Z_div_plus. rewrite (Zdiv_small b y H0). omega. omega. +Qed. + +Lemma Zdiv_Zdiv: + forall a b c, + b > 0 -> c > 0 -> (a / b) / c = a / (b * c). +Proof. + intros. + generalize (Z_div_mod_eq a b H). generalize (Z_mod_lt a b H). intros. + generalize (Z_div_mod_eq (a/b) c H0). generalize (Z_mod_lt (a/b) c H0). intros. + set (q1 := a / b) in *. set (r1 := a mod b) in *. + set (q2 := q1 / c) in *. set (r2 := q1 mod c) in *. + symmetry. apply Zdiv_unique with (r2 * b + r1). + rewrite H2. rewrite H4. ring. + split. + assert (0 <= r2 * b). apply Zmult_le_0_compat. omega. omega. omega. + assert ((r2 + 1) * b <= c * b). + apply Zmult_le_compat_r. omega. omega. + replace ((r2 + 1) * b) with (r2 * b + b) in H5 by ring. + replace (c * b) with (b * c) in H5 by ring. + omega. +Qed. + +Lemma Zmult_le_compat_l_neg : + forall n m p:Z, n >= m -> p <= 0 -> p * n <= p * m. +Proof. + intros. + assert ((-p) * n >= (-p) * m). apply Zmult_ge_compat_l. auto. omega. + replace (p * n) with (- ((-p) * n)) by ring. + replace (p * m) with (- ((-p) * m)) by ring. + omega. +Qed. + +Lemma Zdiv_interval_1: + forall lo hi a b, + lo <= 0 -> hi > 0 -> b > 0 -> + lo * b <= a < hi * b -> + lo <= a/b < hi. +Proof. + intros. + generalize (Z_div_mod_eq a b H1). generalize (Z_mod_lt a b H1). intros. + set (q := a/b) in *. set (r := a mod b) in *. + split. + assert (lo < (q + 1)). + apply Zmult_lt_reg_r with b. omega. + apply Zle_lt_trans with a. omega. + replace ((q + 1) * b) with (b * q + b) by ring. + omega. + omega. + apply Zmult_lt_reg_r with b. omega. + replace (q * b) with (b * q) by ring. + omega. +Qed. + +Lemma Zdiv_interval_2: + forall lo hi a b, + lo <= a <= hi -> lo <= 0 -> hi >= 0 -> b > 0 -> + lo <= a/b <= hi. +Proof. + intros. + assert (lo <= a / b < hi+1). + apply Zdiv_interval_1. omega. omega. auto. + assert (lo * b <= lo * 1). apply Zmult_le_compat_l_neg. omega. omega. + replace (lo * 1) with lo in H3 by ring. + assert ((hi + 1) * 1 <= (hi + 1) * b). apply Zmult_le_compat_l. omega. omega. + replace ((hi + 1) * 1) with (hi + 1) in H4 by ring. + omega. + omega. +Qed. + +Lemma Zmod_recombine: + forall x a b, + a > 0 -> b > 0 -> + x mod (a * b) = ((x/b) mod a) * b + (x mod b). +Proof. + intros. + set (xb := x/b). + apply Zmod_unique with (xb/a). + generalize (Z_div_mod_eq x b H0); fold xb; intro EQ1. + generalize (Z_div_mod_eq xb a H); intro EQ2. + rewrite EQ2 in EQ1. + eapply trans_eq. eexact EQ1. ring. + generalize (Z_mod_lt x b H0). intro. + generalize (Z_mod_lt xb a H). intro. + assert (0 <= xb mod a * b <= a * b - b). + split. apply Zmult_le_0_compat; omega. + replace (a * b - b) with ((a - 1) * b) by ring. + apply Zmult_le_compat; omega. + omega. +Qed. + +(** Properties of divisibility. *) + +Lemma Zdivides_trans: + forall x y z, (x | y) -> (y | z) -> (x | z). +Proof. + intros x y z [a A] [b B]; subst. exists (a*b); ring. +Qed. + +Definition Zdivide_dec: + forall (p q: Z), p > 0 -> { (p|q) } + { ~(p|q) }. +Proof. + intros. destruct (zeq (Zmod q p) 0). + left. exists (q / p). + transitivity (p * (q / p) + (q mod p)). apply Z_div_mod_eq; auto. + transitivity (p * (q / p)). omega. ring. + right; red; intros. elim n. apply Z_div_exact_1; auto. + inv H0. rewrite Z_div_mult; auto. ring. +Defined. +Global Opaque Zdivide_dec. + +Lemma Zdivide_interval: + forall a b c, + 0 < c -> 0 <= a < b -> (c | a) -> (c | b) -> 0 <= a <= b - c. +Proof. + intros. destruct H1 as [x EQ1]. destruct H2 as [y EQ2]. subst. destruct H0. + split. omega. exploit Zmult_lt_reg_r; eauto. intros. + replace (y * c - c) with ((y - 1) * c) by ring. + apply Zmult_le_compat_r; omega. +Qed. + +(** Conversion from [Z] to [nat]. *) + +Definition nat_of_Z: Z -> nat := Z.to_nat. + +Lemma nat_of_Z_of_nat: + forall n, nat_of_Z (Z_of_nat n) = n. +Proof. + exact Nat2Z.id. +Qed. + +Lemma nat_of_Z_max: + forall z, Z_of_nat (nat_of_Z z) = Zmax z 0. +Proof. + intros. unfold Zmax. destruct z; simpl; auto. + change (Z.of_nat (Z.to_nat (Zpos p)) = Zpos p). + apply Z2Nat.id. compute; intuition congruence. +Qed. + +Lemma nat_of_Z_eq: + forall z, z >= 0 -> Z_of_nat (nat_of_Z z) = z. +Proof. + unfold nat_of_Z; intros. apply Z2Nat.id. omega. +Qed. + +Lemma nat_of_Z_neg: + forall n, n <= 0 -> nat_of_Z n = O. +Proof. + destruct n; unfold Zle; simpl; auto. congruence. +Qed. + +Lemma nat_of_Z_plus: + forall p q, + p >= 0 -> q >= 0 -> + nat_of_Z (p + q) = (nat_of_Z p + nat_of_Z q)%nat. +Proof. + unfold nat_of_Z; intros. apply Z2Nat.inj_add; omega. +Qed. + + +(** Alignment: [align n amount] returns the smallest multiple of [amount] + greater than or equal to [n]. *) + +Definition align (n: Z) (amount: Z) := + ((n + amount - 1) / amount) * amount. + +Lemma align_le: forall x y, y > 0 -> x <= align x y. +Proof. + intros. unfold align. + generalize (Z_div_mod_eq (x + y - 1) y H). intro. + replace ((x + y - 1) / y * y) + with ((x + y - 1) - (x + y - 1) mod y). + generalize (Z_mod_lt (x + y - 1) y H). omega. + rewrite Zmult_comm. omega. +Qed. + +Lemma align_divides: forall x y, y > 0 -> (y | align x y). +Proof. + intros. unfold align. apply Zdivide_factor_l. +Qed. + +(** * Definitions and theorems on the data types [option], [sum] and [list] *) + +Set Implicit Arguments. + +(** Comparing option types. *) + +Definition option_eq (A: Type) (eqA: forall (x y: A), {x=y} + {x<>y}): + forall (x y: option A), {x=y} + {x<>y}. +Proof. decide equality. Defined. +Global Opaque option_eq. + +(** Lifting a relation to an option type. *) + +Inductive option_rel (A B: Type) (R: A -> B -> Prop) : option A -> option B -> Prop := + | option_rel_none: option_rel R None None + | option_rel_some: forall x y, R x y -> option_rel R (Some x) (Some y). + +(** Mapping a function over an option type. *) + +Definition option_map (A B: Type) (f: A -> B) (x: option A) : option B := + match x with + | None => None + | Some y => Some (f y) + end. + +(** Mapping a function over a sum type. *) + +Definition sum_left_map (A B C: Type) (f: A -> B) (x: A + C) : B + C := + match x with + | inl y => inl C (f y) + | inr z => inr B z + end. + +(** Properties of [List.nth] (n-th element of a list). *) + +Hint Resolve in_eq in_cons: coqlib. + +Lemma nth_error_in: + forall (A: Type) (n: nat) (l: list A) (x: A), + List.nth_error l n = Some x -> In x l. +Proof. + induction n; simpl. + destruct l; intros. + discriminate. + injection H; intro; subst a. apply in_eq. + destruct l; intros. + discriminate. + apply in_cons. auto. +Qed. +Hint Resolve nth_error_in: coqlib. + +Lemma nth_error_nil: + forall (A: Type) (idx: nat), nth_error (@nil A) idx = None. +Proof. + induction idx; simpl; intros; reflexivity. +Qed. +Hint Resolve nth_error_nil: coqlib. + +(** Compute the length of a list, with result in [Z]. *) + +Fixpoint list_length_z_aux (A: Type) (l: list A) (acc: Z) {struct l}: Z := + match l with + | nil => acc + | hd :: tl => list_length_z_aux tl (Zsucc acc) + end. + +Remark list_length_z_aux_shift: + forall (A: Type) (l: list A) n m, + list_length_z_aux l n = list_length_z_aux l m + (n - m). +Proof. + induction l; intros; simpl. + omega. + replace (n - m) with (Zsucc n - Zsucc m) by omega. auto. +Qed. + +Definition list_length_z (A: Type) (l: list A) : Z := + list_length_z_aux l 0. + +Lemma list_length_z_cons: + forall (A: Type) (hd: A) (tl: list A), + list_length_z (hd :: tl) = list_length_z tl + 1. +Proof. + intros. unfold list_length_z. simpl. + rewrite (list_length_z_aux_shift tl 1 0). omega. +Qed. + +Lemma list_length_z_pos: + forall (A: Type) (l: list A), + list_length_z l >= 0. +Proof. + induction l; simpl. unfold list_length_z; simpl. omega. + rewrite list_length_z_cons. omega. +Qed. + +Lemma list_length_z_map: + forall (A B: Type) (f: A -> B) (l: list A), + list_length_z (map f l) = list_length_z l. +Proof. + induction l. reflexivity. simpl. repeat rewrite list_length_z_cons. congruence. +Qed. + +(** Extract the n-th element of a list, as [List.nth_error] does, + but the index [n] is of type [Z]. *) + +Fixpoint list_nth_z (A: Type) (l: list A) (n: Z) {struct l}: option A := + match l with + | nil => None + | hd :: tl => if zeq n 0 then Some hd else list_nth_z tl (Zpred n) + end. + +Lemma list_nth_z_in: + forall (A: Type) (l: list A) n x, + list_nth_z l n = Some x -> In x l. +Proof. + induction l; simpl; intros. + congruence. + destruct (zeq n 0). left; congruence. right; eauto. +Qed. + +Lemma list_nth_z_map: + forall (A B: Type) (f: A -> B) (l: list A) n, + list_nth_z (List.map f l) n = option_map f (list_nth_z l n). +Proof. + induction l; simpl; intros. + auto. + destruct (zeq n 0). auto. eauto. +Qed. + +Lemma list_nth_z_range: + forall (A: Type) (l: list A) n x, + list_nth_z l n = Some x -> 0 <= n < list_length_z l. +Proof. + induction l; simpl; intros. + discriminate. + rewrite list_length_z_cons. destruct (zeq n 0). + generalize (list_length_z_pos l); omega. + exploit IHl; eauto. omega. +Qed. + +(** Properties of [List.incl] (list inclusion). *) + +Lemma incl_cons_inv: + forall (A: Type) (a: A) (b c: list A), + incl (a :: b) c -> incl b c. +Proof. + unfold incl; intros. apply H. apply in_cons. auto. +Qed. +Hint Resolve incl_cons_inv: coqlib. + +Lemma incl_app_inv_l: + forall (A: Type) (l1 l2 m: list A), + incl (l1 ++ l2) m -> incl l1 m. +Proof. + unfold incl; intros. apply H. apply in_or_app. left; assumption. +Qed. + +Lemma incl_app_inv_r: + forall (A: Type) (l1 l2 m: list A), + incl (l1 ++ l2) m -> incl l2 m. +Proof. + unfold incl; intros. apply H. apply in_or_app. right; assumption. +Qed. + +Hint Resolve incl_tl incl_refl incl_app_inv_l incl_app_inv_r: coqlib. + +Lemma incl_same_head: + forall (A: Type) (x: A) (l1 l2: list A), + incl l1 l2 -> incl (x::l1) (x::l2). +Proof. + intros; red; simpl; intros. intuition. +Qed. + +(** Properties of [List.map] (mapping a function over a list). *) + +Lemma list_map_exten: + forall (A B: Type) (f f': A -> B) (l: list A), + (forall x, In x l -> f x = f' x) -> + List.map f' l = List.map f l. +Proof. + induction l; simpl; intros. + reflexivity. + rewrite <- H. rewrite IHl. reflexivity. + intros. apply H. tauto. + tauto. +Qed. + +Lemma list_map_compose: + forall (A B C: Type) (f: A -> B) (g: B -> C) (l: list A), + List.map g (List.map f l) = List.map (fun x => g(f x)) l. +Proof. + induction l; simpl. reflexivity. rewrite IHl; reflexivity. +Qed. + +Lemma list_map_identity: + forall (A: Type) (l: list A), + List.map (fun (x:A) => x) l = l. +Proof. + induction l; simpl; congruence. +Qed. + +Lemma list_map_nth: + forall (A B: Type) (f: A -> B) (l: list A) (n: nat), + nth_error (List.map f l) n = option_map f (nth_error l n). +Proof. + induction l; simpl; intros. + repeat rewrite nth_error_nil. reflexivity. + destruct n; simpl. reflexivity. auto. +Qed. + +Lemma list_length_map: + forall (A B: Type) (f: A -> B) (l: list A), + List.length (List.map f l) = List.length l. +Proof. + induction l; simpl; congruence. +Qed. + +Lemma list_in_map_inv: + forall (A B: Type) (f: A -> B) (l: list A) (y: B), + In y (List.map f l) -> exists x:A, y = f x /\ In x l. +Proof. + induction l; simpl; intros. + contradiction. + elim H; intro. + exists a; intuition auto. + generalize (IHl y H0). intros [x [EQ IN]]. + exists x; tauto. +Qed. + +Lemma list_append_map: + forall (A B: Type) (f: A -> B) (l1 l2: list A), + List.map f (l1 ++ l2) = List.map f l1 ++ List.map f l2. +Proof. + induction l1; simpl; intros. + auto. rewrite IHl1. auto. +Qed. + +Lemma list_append_map_inv: + forall (A B: Type) (f: A -> B) (m1 m2: list B) (l: list A), + List.map f l = m1 ++ m2 -> + exists l1, exists l2, List.map f l1 = m1 /\ List.map f l2 = m2 /\ l = l1 ++ l2. +Proof. + induction m1; simpl; intros. + exists (@nil A); exists l; auto. + destruct l; simpl in H; inv H. + exploit IHm1; eauto. intros [l1 [l2 [P [Q R]]]]. subst l. + exists (a0 :: l1); exists l2; intuition. simpl; congruence. +Qed. + +(** Folding a function over a list *) + +Section LIST_FOLD. + +Variables A B: Type. +Variable f: A -> B -> B. + +(** This is exactly [List.fold_left] from Coq's standard library, + with [f] taking arguments in a different order. *) + +Fixpoint list_fold_left (accu: B) (l: list A) : B := + match l with nil => accu | x :: l' => list_fold_left (f x accu) l' end. + +(** This is exactly [List.fold_right] from Coq's standard library, + except that it runs in constant stack space. *) + +Definition list_fold_right (l: list A) (base: B) : B := + list_fold_left base (List.rev' l). + +Remark list_fold_left_app: + forall l1 l2 accu, + list_fold_left accu (l1 ++ l2) = list_fold_left (list_fold_left accu l1) l2. +Proof. + induction l1; simpl; intros. + auto. + rewrite IHl1. auto. +Qed. + +Lemma list_fold_right_eq: + forall l base, + list_fold_right l base = + match l with nil => base | x :: l' => f x (list_fold_right l' base) end. +Proof. + unfold list_fold_right; intros. + destruct l. + auto. + unfold rev'. rewrite <- ! rev_alt. simpl. + rewrite list_fold_left_app. simpl. auto. +Qed. + +Lemma list_fold_right_spec: + forall l base, list_fold_right l base = List.fold_right f base l. +Proof. + induction l; simpl; intros; rewrite list_fold_right_eq; congruence. +Qed. + +End LIST_FOLD. + +(** Properties of list membership. *) + +Lemma in_cns: + forall (A: Type) (x y: A) (l: list A), In x (y :: l) <-> y = x \/ In x l. +Proof. + intros. simpl. tauto. +Qed. + +Lemma in_app: + forall (A: Type) (x: A) (l1 l2: list A), In x (l1 ++ l2) <-> In x l1 \/ In x l2. +Proof. + intros. split; intro. apply in_app_or. auto. apply in_or_app. auto. +Qed. + +Lemma list_in_insert: + forall (A: Type) (x: A) (l1 l2: list A) (y: A), + In x (l1 ++ l2) -> In x (l1 ++ y :: l2). +Proof. + intros. apply in_or_app; simpl. elim (in_app_or _ _ _ H); intro; auto. +Qed. + +(** [list_disjoint l1 l2] holds iff [l1] and [l2] have no elements + in common. *) + +Definition list_disjoint (A: Type) (l1 l2: list A) : Prop := + forall (x y: A), In x l1 -> In y l2 -> x <> y. + +Lemma list_disjoint_cons_l: + forall (A: Type) (a: A) (l1 l2: list A), + list_disjoint l1 l2 -> ~In a l2 -> list_disjoint (a :: l1) l2. +Proof. + unfold list_disjoint; simpl; intros. destruct H1. congruence. apply H; auto. +Qed. + +Lemma list_disjoint_cons_r: + forall (A: Type) (a: A) (l1 l2: list A), + list_disjoint l1 l2 -> ~In a l1 -> list_disjoint l1 (a :: l2). +Proof. + unfold list_disjoint; simpl; intros. destruct H2. congruence. apply H; auto. +Qed. + +Lemma list_disjoint_cons_left: + forall (A: Type) (a: A) (l1 l2: list A), + list_disjoint (a :: l1) l2 -> list_disjoint l1 l2. +Proof. + unfold list_disjoint; simpl; intros. apply H; tauto. +Qed. + +Lemma list_disjoint_cons_right: + forall (A: Type) (a: A) (l1 l2: list A), + list_disjoint l1 (a :: l2) -> list_disjoint l1 l2. +Proof. + unfold list_disjoint; simpl; intros. apply H; tauto. +Qed. + +Lemma list_disjoint_notin: + forall (A: Type) (l1 l2: list A) (a: A), + list_disjoint l1 l2 -> In a l1 -> ~(In a l2). +Proof. + unfold list_disjoint; intros; red; intros. + apply H with a a; auto. +Qed. + +Lemma list_disjoint_sym: + forall (A: Type) (l1 l2: list A), + list_disjoint l1 l2 -> list_disjoint l2 l1. +Proof. + unfold list_disjoint; intros. + apply sym_not_equal. apply H; auto. +Qed. + +Lemma list_disjoint_dec: + forall (A: Type) (eqA_dec: forall (x y: A), {x=y} + {x<>y}) (l1 l2: list A), + {list_disjoint l1 l2} + {~list_disjoint l1 l2}. +Proof. + induction l1; intros. + left; red; intros. elim H. + case (In_dec eqA_dec a l2); intro. + right; red; intro. apply (H a a); auto with coqlib. + case (IHl1 l2); intro. + left; red; intros. elim H; intro. + red; intro; subst a y. contradiction. + apply l; auto. + right; red; intros. elim n0. eapply list_disjoint_cons_left; eauto. +Defined. + +(** [list_equiv l1 l2] holds iff the lists [l1] and [l2] contain the same elements. *) + +Definition list_equiv (A : Type) (l1 l2: list A) : Prop := + forall x, In x l1 <-> In x l2. + +(** [list_norepet l] holds iff the list [l] contains no repetitions, + i.e. no element occurs twice. *) + +Inductive list_norepet (A: Type) : list A -> Prop := + | list_norepet_nil: + list_norepet nil + | list_norepet_cons: + forall hd tl, + ~(In hd tl) -> list_norepet tl -> list_norepet (hd :: tl). + +Lemma list_norepet_dec: + forall (A: Type) (eqA_dec: forall (x y: A), {x=y} + {x<>y}) (l: list A), + {list_norepet l} + {~list_norepet l}. +Proof. + induction l. + left; constructor. + destruct IHl. + case (In_dec eqA_dec a l); intro. + right. red; intro. inversion H. contradiction. + left. constructor; auto. + right. red; intro. inversion H. contradiction. +Defined. + +Lemma list_map_norepet: + forall (A B: Type) (f: A -> B) (l: list A), + list_norepet l -> + (forall x y, In x l -> In y l -> x <> y -> f x <> f y) -> + list_norepet (List.map f l). +Proof. + induction 1; simpl; intros. + constructor. + constructor. + red; intro. generalize (list_in_map_inv f _ _ H2). + intros [x [EQ IN]]. generalize EQ. change (f hd <> f x). + apply H1. tauto. tauto. + red; intro; subst x. contradiction. + apply IHlist_norepet. intros. apply H1. tauto. tauto. auto. +Qed. + +Remark list_norepet_append_commut: + forall (A: Type) (a b: list A), + list_norepet (a ++ b) -> list_norepet (b ++ a). +Proof. + intro A. + assert (forall (x: A) (b: list A) (a: list A), + list_norepet (a ++ b) -> ~(In x a) -> ~(In x b) -> + list_norepet (a ++ x :: b)). + induction a; simpl; intros. + constructor; auto. + inversion H. constructor. red; intro. + elim (in_app_or _ _ _ H6); intro. + elim H4. apply in_or_app. tauto. + elim H7; intro. subst a. elim H0. left. auto. + elim H4. apply in_or_app. tauto. + auto. + induction a; simpl; intros. + rewrite <- app_nil_end. auto. + inversion H0. apply H. auto. + red; intro; elim H3. apply in_or_app. tauto. + red; intro; elim H3. apply in_or_app. tauto. +Qed. + +Lemma list_norepet_app: + forall (A: Type) (l1 l2: list A), + list_norepet (l1 ++ l2) <-> + list_norepet l1 /\ list_norepet l2 /\ list_disjoint l1 l2. +Proof. + induction l1; simpl; intros; split; intros. + intuition. constructor. red;simpl;auto. + tauto. + inversion H; subst. rewrite IHl1 in H3. rewrite in_app in H2. + intuition. + constructor; auto. red; intros. elim H2; intro. congruence. auto. + destruct H as [B [C D]]. inversion B; subst. + constructor. rewrite in_app. intuition. elim (D a a); auto. apply in_eq. + rewrite IHl1. intuition. red; intros. apply D; auto. apply in_cons; auto. +Qed. + +Lemma list_norepet_append: + forall (A: Type) (l1 l2: list A), + list_norepet l1 -> list_norepet l2 -> list_disjoint l1 l2 -> + list_norepet (l1 ++ l2). +Proof. + generalize list_norepet_app; firstorder. +Qed. + +Lemma list_norepet_append_right: + forall (A: Type) (l1 l2: list A), + list_norepet (l1 ++ l2) -> list_norepet l2. +Proof. + generalize list_norepet_app; firstorder. +Qed. + +Lemma list_norepet_append_left: + forall (A: Type) (l1 l2: list A), + list_norepet (l1 ++ l2) -> list_norepet l1. +Proof. + generalize list_norepet_app; firstorder. +Qed. + +(** [is_tail l1 l2] holds iff [l2] is of the form [l ++ l1] for some [l]. *) + +Inductive is_tail (A: Type): list A -> list A -> Prop := + | is_tail_refl: + forall c, is_tail c c + | is_tail_cons: + forall i c1 c2, is_tail c1 c2 -> is_tail c1 (i :: c2). + +Lemma is_tail_in: + forall (A: Type) (i: A) c1 c2, is_tail (i :: c1) c2 -> In i c2. +Proof. + induction c2; simpl; intros. + inversion H. + inversion H. tauto. right; auto. +Qed. + +Lemma is_tail_cons_left: + forall (A: Type) (i: A) c1 c2, is_tail (i :: c1) c2 -> is_tail c1 c2. +Proof. + induction c2; intros; inversion H. + constructor. constructor. constructor. auto. +Qed. + +Hint Resolve is_tail_refl is_tail_cons is_tail_in is_tail_cons_left: coqlib. + +Lemma is_tail_incl: + forall (A: Type) (l1 l2: list A), is_tail l1 l2 -> incl l1 l2. +Proof. + induction 1; eauto with coqlib. +Qed. + +Lemma is_tail_trans: + forall (A: Type) (l1 l2: list A), + is_tail l1 l2 -> forall (l3: list A), is_tail l2 l3 -> is_tail l1 l3. +Proof. + induction 1; intros. auto. apply IHis_tail. eapply is_tail_cons_left; eauto. +Qed. + +(** [list_forall2 P [x1 ... xN] [y1 ... yM]] holds iff [N = M] and + [P xi yi] holds for all [i]. *) + +Section FORALL2. + +Variable A: Type. +Variable B: Type. +Variable P: A -> B -> Prop. + +Inductive list_forall2: list A -> list B -> Prop := + | list_forall2_nil: + list_forall2 nil nil + | list_forall2_cons: + forall a1 al b1 bl, + P a1 b1 -> + list_forall2 al bl -> + list_forall2 (a1 :: al) (b1 :: bl). + +Lemma list_forall2_app: + forall a2 b2 a1 b1, + list_forall2 a1 b1 -> list_forall2 a2 b2 -> + list_forall2 (a1 ++ a2) (b1 ++ b2). +Proof. + induction 1; intros; simpl. auto. constructor; auto. +Qed. + +Lemma list_forall2_length: + forall l1 l2, + list_forall2 l1 l2 -> length l1 = length l2. +Proof. + induction 1; simpl; congruence. +Qed. + +Lemma list_forall2_in_left: + forall x1 l1 l2, + list_forall2 l1 l2 -> In x1 l1 -> exists x2, In x2 l2 /\ P x1 x2. +Proof. + induction 1; simpl; intros. contradiction. destruct H1. + subst; exists b1; auto. + exploit IHlist_forall2; eauto. intros (x2 & U & V); exists x2; auto. +Qed. + +Lemma list_forall2_in_right: + forall x2 l1 l2, + list_forall2 l1 l2 -> In x2 l2 -> exists x1, In x1 l1 /\ P x1 x2. +Proof. + induction 1; simpl; intros. contradiction. destruct H1. + subst; exists a1; auto. + exploit IHlist_forall2; eauto. intros (x1 & U & V); exists x1; auto. +Qed. + +End FORALL2. + +Lemma list_forall2_imply: + forall (A B: Type) (P1: A -> B -> Prop) (l1: list A) (l2: list B), + list_forall2 P1 l1 l2 -> + forall (P2: A -> B -> Prop), + (forall v1 v2, In v1 l1 -> In v2 l2 -> P1 v1 v2 -> P2 v1 v2) -> + list_forall2 P2 l1 l2. +Proof. + induction 1; intros. + constructor. + constructor. auto with coqlib. apply IHlist_forall2; auto. + intros. auto with coqlib. +Qed. + +(** Dropping the first N elements of a list. *) + +Fixpoint list_drop (A: Type) (n: nat) (x: list A) {struct n} : list A := + match n with + | O => x + | S n' => match x with nil => nil | hd :: tl => list_drop n' tl end + end. + +Lemma list_drop_incl: + forall (A: Type) (x: A) n (l: list A), In x (list_drop n l) -> In x l. +Proof. + induction n; simpl; intros. auto. + destruct l; auto with coqlib. +Qed. + +Lemma list_drop_norepet: + forall (A: Type) n (l: list A), list_norepet l -> list_norepet (list_drop n l). +Proof. + induction n; simpl; intros. auto. + inv H. constructor. auto. +Qed. + +Lemma list_map_drop: + forall (A B: Type) (f: A -> B) n (l: list A), + list_drop n (map f l) = map f (list_drop n l). +Proof. + induction n; simpl; intros. auto. + destruct l; simpl; auto. +Qed. + +(** A list of [n] elements, all equal to [x]. *) + +Fixpoint list_repeat {A: Type} (n: nat) (x: A) {struct n} := + match n with + | O => nil + | S m => x :: list_repeat m x + end. + +Lemma length_list_repeat: + forall (A: Type) n (x: A), length (list_repeat n x) = n. +Proof. + induction n; simpl; intros. auto. decEq; auto. +Qed. + +Lemma in_list_repeat: + forall (A: Type) n (x: A) y, In y (list_repeat n x) -> y = x. +Proof. + induction n; simpl; intros. elim H. destruct H; auto. +Qed. + +(** * Definitions and theorems over boolean types *) + +Definition proj_sumbool {P Q: Prop} (a: {P} + {Q}) : bool := + if a then true else false. + +Coercion proj_sumbool: sumbool >-> bool. + +Lemma proj_sumbool_true: + forall (P Q: Prop) (a: {P}+{Q}), proj_sumbool a = true -> P. +Proof. + intros P Q a. destruct a; simpl. auto. congruence. +Qed. + +Lemma proj_sumbool_is_true: + forall (P: Prop) (a: {P}+{~P}), P -> proj_sumbool a = true. +Proof. + intros. unfold proj_sumbool. destruct a. auto. contradiction. +Qed. + +Ltac InvBooleans := + match goal with + | [ H: _ && _ = true |- _ ] => + destruct (andb_prop _ _ H); clear H; InvBooleans + | [ H: _ || _ = false |- _ ] => + destruct (orb_false_elim _ _ H); clear H; InvBooleans + | [ H: proj_sumbool ?x = true |- _ ] => + generalize (proj_sumbool_true _ H); clear H; intro; InvBooleans + | _ => idtac + end. + +Section DECIDABLE_EQUALITY. + +Variable A: Type. +Variable dec_eq: forall (x y: A), {x=y} + {x<>y}. +Variable B: Type. + +Lemma dec_eq_true: + forall (x: A) (ifso ifnot: B), + (if dec_eq x x then ifso else ifnot) = ifso. +Proof. + intros. destruct (dec_eq x x). auto. congruence. +Qed. + +Lemma dec_eq_false: + forall (x y: A) (ifso ifnot: B), + x <> y -> (if dec_eq x y then ifso else ifnot) = ifnot. +Proof. + intros. destruct (dec_eq x y). congruence. auto. +Qed. + +Lemma dec_eq_sym: + forall (x y: A) (ifso ifnot: B), + (if dec_eq x y then ifso else ifnot) = + (if dec_eq y x then ifso else ifnot). +Proof. + intros. destruct (dec_eq x y). + subst y. rewrite dec_eq_true. auto. + rewrite dec_eq_false; auto. +Qed. + +End DECIDABLE_EQUALITY. + +Section DECIDABLE_PREDICATE. + +Variable P: Prop. +Variable dec: {P} + {~P}. +Variable A: Type. + +Lemma pred_dec_true: + forall (a b: A), P -> (if dec then a else b) = a. +Proof. + intros. destruct dec. auto. contradiction. +Qed. + +Lemma pred_dec_false: + forall (a b: A), ~P -> (if dec then a else b) = b. +Proof. + intros. destruct dec. contradiction. auto. +Qed. + +End DECIDABLE_PREDICATE. + +(** * Well-founded orderings *) + +Require Import Relations. + +(** A non-dependent version of lexicographic ordering. *) + +Section LEX_ORDER. + +Variable A: Type. +Variable B: Type. +Variable ordA: A -> A -> Prop. +Variable ordB: B -> B -> Prop. + +Inductive lex_ord: A*B -> A*B -> Prop := + | lex_ord_left: forall a1 b1 a2 b2, + ordA a1 a2 -> lex_ord (a1,b1) (a2,b2) + | lex_ord_right: forall a b1 b2, + ordB b1 b2 -> lex_ord (a,b1) (a,b2). + +Lemma wf_lex_ord: + well_founded ordA -> well_founded ordB -> well_founded lex_ord. +Proof. + intros Awf Bwf. + assert (forall a, Acc ordA a -> forall b, Acc ordB b -> Acc lex_ord (a, b)). + induction 1. induction 1. constructor; intros. inv H3. + apply H0. auto. apply Bwf. + apply H2; auto. + red; intros. destruct a as [a b]. apply H; auto. +Qed. + +Lemma transitive_lex_ord: + transitive _ ordA -> transitive _ ordB -> transitive _ lex_ord. +Proof. + intros trA trB; red; intros. + inv H; inv H0. + left; eapply trA; eauto. + left; auto. + left; auto. + right; eapply trB; eauto. +Qed. + +End LEX_ORDER. + +(** * Nonempty lists *) + +Inductive nlist (A: Type) : Type := + | nbase: A -> nlist A + | ncons: A -> nlist A -> nlist A. + +Definition nfirst {A: Type} (l: nlist A) := + match l with nbase a => a | ncons a l' => a end. + +Fixpoint nlast {A: Type} (l: nlist A) := + match l with nbase a => a | ncons a l' => nlast l' end. + +Fixpoint nIn {A: Type} (x: A) (l: nlist A) : Prop := + match l with + | nbase a => a = x + | ncons a l => a = x \/ nIn x l + end. + +Inductive nlist_forall2 {A B: Type} (R: A -> B -> Prop) : nlist A -> nlist B -> Prop := + | nbase_forall2: forall a b, R a b -> nlist_forall2 R (nbase a) (nbase b) + | ncons_forall2: forall a l b m, R a b -> nlist_forall2 R l m -> nlist_forall2 R (ncons a l) (ncons b m). + +Lemma nlist_forall2_imply: + forall (A B: Type) (P1: A -> B -> Prop) (l1: nlist A) (l2: nlist B), + nlist_forall2 P1 l1 l2 -> + forall (P2: A -> B -> Prop), + (forall v1 v2, nIn v1 l1 -> nIn v2 l2 -> P1 v1 v2 -> P2 v1 v2) -> + nlist_forall2 P2 l1 l2. +Proof. + induction 1; simpl; intros; constructor; auto. +Qed. diff --git a/src/cclib/Errors.v b/src/cclib/Errors.v new file mode 100755 index 0000000..effd951 --- /dev/null +++ b/src/cclib/Errors.v @@ -0,0 +1,193 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Error reporting and the error monad. *) + +Require Import String. +Require Import Coqlib. + +Close Scope string_scope. + +Set Implicit Arguments. + +(** * Representation of error messages. *) + +(** Compile-time errors produce an error message, represented in Coq + as a list of either substrings or positive numbers encoding + a source-level identifier (see module AST). *) + +Inductive errcode: Type := + | MSG: string -> errcode + | CTX: positive -> errcode (* a top-level identifier *) + | POS: positive -> errcode. (* a positive integer, e.g. a PC *) + +Definition errmsg: Type := list errcode. + +Definition msg (s: string) : errmsg := MSG s :: nil. + +(** * The error monad *) + +(** Compilation functions that can fail have return type [res A]. + The return value is either [OK res] to indicate success, + or [Error msg] to indicate failure. *) + +Inductive res (A: Type) : Type := +| OK: A -> res A +| Error: errmsg -> res A. + +Arguments Error [A]. + +(** To automate the propagation of errors, we use a monadic style + with the following [bind] operation. *) + +Definition bind (A B: Type) (f: res A) (g: A -> res B) : res B := + match f with + | OK x => g x + | Error msg => Error msg + end. + +Definition bind2 (A B C: Type) (f: res (A * B)) (g: A -> B -> res C) : res C := + match f with + | OK (x, y) => g x y + | Error msg => Error msg + end. + +(** The [do] notation, inspired by Haskell's, keeps the code readable. *) + +Notation "'do' X <- A ; B" := (bind A (fun X => B)) + (at level 200, X ident, A at level 100, B at level 200) + : error_monad_scope. + +Notation "'do' ( X , Y ) <- A ; B" := (bind2 A (fun X Y => B)) + (at level 200, X ident, Y ident, A at level 100, B at level 200) + : error_monad_scope. + +Remark bind_inversion: + forall (A B: Type) (f: res A) (g: A -> res B) (y: B), + bind f g = OK y -> + exists x, f = OK x /\ g x = OK y. +Proof. + intros until y. destruct f; simpl; intros. + exists a; auto. + discriminate. +Qed. + +Remark bind2_inversion: + forall (A B C: Type) (f: res (A*B)) (g: A -> B -> res C) (z: C), + bind2 f g = OK z -> + exists x, exists y, f = OK (x, y) /\ g x y = OK z. +Proof. + intros until z. destruct f; simpl. + destruct p; simpl; intros. exists a; exists b; auto. + intros; discriminate. +Qed. + +(** Assertions *) + +Definition assertion_failed {A: Type} : res A := Error(msg "Assertion failed"). + +Notation "'assertion' A ; B" := (if A then B else assertion_failed) + (at level 200, A at level 100, B at level 200) + : error_monad_scope. + +(** This is the familiar monadic map iterator. *) + +Local Open Scope error_monad_scope. + +Fixpoint mmap (A B: Type) (f: A -> res B) (l: list A) {struct l} : res (list B) := + match l with + | nil => OK nil + | hd :: tl => do hd' <- f hd; do tl' <- mmap f tl; OK (hd' :: tl') + end. + +Remark mmap_inversion: + forall (A B: Type) (f: A -> res B) (l: list A) (l': list B), + mmap f l = OK l' -> + list_forall2 (fun x y => f x = OK y) l l'. +Proof. + induction l; simpl; intros. + inversion_clear H. constructor. + destruct (bind_inversion _ _ H) as [hd' [P Q]]. + destruct (bind_inversion _ _ Q) as [tl' [R S]]. + inversion_clear S. + constructor. auto. auto. +Qed. + +(** * Reasoning over monadic computations *) + +(** The [monadInv H] tactic below simplifies hypotheses of the form +<< + H: (do x <- a; b) = OK res +>> + By definition of the bind operation, both computations [a] and + [b] must succeed for their composition to succeed. The tactic + therefore generates the following hypotheses: + + x: ... + H1: a = OK x + H2: b x = OK res +*) + +Ltac monadInv1 H := + match type of H with + | (OK _ = OK _) => + inversion H; clear H; try subst + | (Error _ = OK _) => + discriminate + | (bind ?F ?G = OK ?X) => + let x := fresh "x" in ( + let EQ1 := fresh "EQ" in ( + let EQ2 := fresh "EQ" in ( + destruct (bind_inversion F G H) as [x [EQ1 EQ2]]; + clear H; + try (monadInv1 EQ2)))) + | (bind2 ?F ?G = OK ?X) => + let x1 := fresh "x" in ( + let x2 := fresh "x" in ( + let EQ1 := fresh "EQ" in ( + let EQ2 := fresh "EQ" in ( + destruct (bind2_inversion F G H) as [x1 [x2 [EQ1 EQ2]]]; + clear H; + try (monadInv1 EQ2))))) + | (match ?X with left _ => _ | right _ => assertion_failed end = OK _) => + destruct X; [try (monadInv1 H) | discriminate] + | (match (negb ?X) with true => _ | false => assertion_failed end = OK _) => + destruct X as [] eqn:?; [discriminate | try (monadInv1 H)] + | (match ?X with true => _ | false => assertion_failed end = OK _) => + destruct X as [] eqn:?; [try (monadInv1 H) | discriminate] + | (mmap ?F ?L = OK ?M) => + generalize (mmap_inversion F L H); intro + end. + +Ltac monadInv H := + monadInv1 H || + match type of H with + | (?F _ _ _ _ _ _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ = OK _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + end. diff --git a/src/cclib/Integers.v b/src/cclib/Integers.v new file mode 100755 index 0000000..b0e35c1 --- /dev/null +++ b/src/cclib/Integers.v @@ -0,0 +1,5266 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Formalizations of machine integers modulo $2^N$ #2N#. *) + +Require Import Eqdep_dec Zquot Zwf. +Require Import cclib.Coqlib. + +(** * Comparisons *) + +Inductive comparison : Type := + | Ceq : comparison (**r same *) + | Cne : comparison (**r different *) + | Clt : comparison (**r less than *) + | Cle : comparison (**r less than or equal *) + | Cgt : comparison (**r greater than *) + | Cge : comparison. (**r greater than or equal *) + +Definition negate_comparison (c: comparison): comparison := + match c with + | Ceq => Cne + | Cne => Ceq + | Clt => Cge + | Cle => Cgt + | Cgt => Cle + | Cge => Clt + end. + +Definition swap_comparison (c: comparison): comparison := + match c with + | Ceq => Ceq + | Cne => Cne + | Clt => Cgt + | Cle => Cge + | Cgt => Clt + | Cge => Cle + end. + +(** * Parameterization by the word size, in bits. *) + +Module Type WORDSIZE. + Parameter wordsize: nat. + Axiom wordsize_not_zero: wordsize <> 0%nat. +End WORDSIZE. + +(* To avoid useless definitions of inductors in extracted code. *) +Local Unset Elimination Schemes. +Local Unset Case Analysis Schemes. + +Module Make(WS: WORDSIZE). + +Definition wordsize: nat := WS.wordsize. +Definition zwordsize: Z := Z_of_nat wordsize. +Definition modulus : Z := two_power_nat wordsize. +Definition half_modulus : Z := modulus / 2. +Definition max_unsigned : Z := modulus - 1. +Definition max_signed : Z := half_modulus - 1. +Definition min_signed : Z := - half_modulus. + +Remark wordsize_pos: zwordsize > 0. +Proof. + unfold zwordsize, wordsize. generalize WS.wordsize_not_zero. omega. +Qed. + +Remark modulus_power: modulus = two_p zwordsize. +Proof. + unfold modulus. apply two_power_nat_two_p. +Qed. + +Remark modulus_pos: modulus > 0. +Proof. + rewrite modulus_power. apply two_p_gt_ZERO. generalize wordsize_pos; omega. +Qed. + +(** * Representation of machine integers *) + +(** A machine integer (type [int]) is represented as a Coq arbitrary-precision + integer (type [Z]) plus a proof that it is in the range 0 (included) to + [modulus] (excluded). *) + +Record int: Type := mkint { intval: Z; intrange: -1 < intval < modulus }. + +(** Fast normalization modulo [2^wordsize] *) + +Fixpoint P_mod_two_p (p: positive) (n: nat) {struct n} : Z := + match n with + | O => 0 + | S m => + match p with + | xH => 1 + | xO q => Z.double (P_mod_two_p q m) + | xI q => Z.succ_double (P_mod_two_p q m) + end + end. + +Definition Z_mod_modulus (x: Z) : Z := + match x with + | Z0 => 0 + | Zpos p => P_mod_two_p p wordsize + | Zneg p => let r := P_mod_two_p p wordsize in if zeq r 0 then 0 else modulus - r + end. + +Lemma P_mod_two_p_range: + forall n p, 0 <= P_mod_two_p p n < two_power_nat n. +Proof. + induction n; simpl; intros. + - rewrite two_power_nat_O. omega. + - rewrite two_power_nat_S. destruct p. + + generalize (IHn p). rewrite Z.succ_double_spec. omega. + + generalize (IHn p). rewrite Z.double_spec. omega. + + generalize (two_power_nat_pos n). omega. +Qed. + +Lemma P_mod_two_p_eq: + forall n p, P_mod_two_p p n = (Zpos p) mod (two_power_nat n). +Proof. + assert (forall n p, exists y, Zpos p = y * two_power_nat n + P_mod_two_p p n). + { + induction n; simpl; intros. + - rewrite two_power_nat_O. exists (Zpos p). ring. + - rewrite two_power_nat_S. destruct p. + + destruct (IHn p) as [y EQ]. exists y. + change (Zpos p~1) with (2 * Zpos p + 1). rewrite EQ. + rewrite Z.succ_double_spec. ring. + + destruct (IHn p) as [y EQ]. exists y. + change (Zpos p~0) with (2 * Zpos p). rewrite EQ. + rewrite (Z.double_spec (P_mod_two_p p n)). ring. + + exists 0; omega. + } + intros. + destruct (H n p) as [y EQ]. + symmetry. apply Zmod_unique with y. auto. apply P_mod_two_p_range. +Qed. + +Lemma Z_mod_modulus_range: + forall x, 0 <= Z_mod_modulus x < modulus. +Proof. + intros; unfold Z_mod_modulus. + destruct x. + - generalize modulus_pos; omega. + - apply P_mod_two_p_range. + - set (r := P_mod_two_p p wordsize). + assert (0 <= r < modulus) by apply P_mod_two_p_range. + destruct (zeq r 0). + + generalize modulus_pos; omega. + + omega. +Qed. + +Lemma Z_mod_modulus_range': + forall x, -1 < Z_mod_modulus x < modulus. +Proof. + intros. generalize (Z_mod_modulus_range x); omega. +Qed. + +Lemma Z_mod_modulus_eq: + forall x, Z_mod_modulus x = x mod modulus. +Proof. + intros. unfold Z_mod_modulus. destruct x. + - rewrite Zmod_0_l. auto. + - apply P_mod_two_p_eq. + - generalize (P_mod_two_p_range wordsize p) (P_mod_two_p_eq wordsize p). + fold modulus. intros A B. + exploit (Z_div_mod_eq (Zpos p) modulus). apply modulus_pos. intros C. + set (q := Zpos p / modulus) in *. + set (r := P_mod_two_p p wordsize) in *. + rewrite <- B in C. + change (Z.neg p) with (- (Z.pos p)). destruct (zeq r 0). + + symmetry. apply Zmod_unique with (-q). rewrite C; rewrite e. ring. + generalize modulus_pos; omega. + + symmetry. apply Zmod_unique with (-q - 1). rewrite C. ring. + omega. +Qed. + +(** The [unsigned] and [signed] functions return the Coq integer corresponding + to the given machine integer, interpreted as unsigned or signed + respectively. *) + +Definition unsigned (n: int) : Z := intval n. + +Definition signed (n: int) : Z := + let x := unsigned n in + if zlt x half_modulus then x else x - modulus. + +(** Conversely, [repr] takes a Coq integer and returns the corresponding + machine integer. The argument is treated modulo [modulus]. *) + +Definition repr (x: Z) : int := + mkint (Z_mod_modulus x) (Z_mod_modulus_range' x). + +Definition zero := repr 0. +Definition one := repr 1. +Definition mone := repr (-1). +Definition iwordsize := repr zwordsize. + +Lemma mkint_eq: + forall x y Px Py, x = y -> mkint x Px = mkint y Py. +Proof. + intros. subst y. + assert (forall (n m: Z) (P1 P2: n < m), P1 = P2). + { + unfold Zlt; intros. + apply eq_proofs_unicity. + intros c1 c2. destruct c1; destruct c2; (left; reflexivity) || (right; congruence). + } + destruct Px as [Px1 Px2]. destruct Py as [Py1 Py2]. + rewrite (H _ _ Px1 Py1). + rewrite (H _ _ Px2 Py2). + reflexivity. +Qed. + +Lemma eq_dec: forall (x y: int), {x = y} + {x <> y}. +Proof. + intros. destruct x; destruct y. destruct (zeq intval0 intval1). + left. apply mkint_eq. auto. + right. red; intro. injection H. exact n. +Defined. + +(* Repr is injective for integers which are already in bound. *) + +Lemma repr_injective : forall x y, + (-1 < x < modulus) -> + (-1 < y < modulus) -> + repr x = repr y -> + x = y. +Proof. + intros x y xbound ybound H. + assert (intval (repr x) = intval (repr y)) as H1. + rewrite H; reflexivity. + clear H. + unfold repr in H1; simpl in H1. + repeat rewrite Z_mod_modulus_eq in H1. + repeat rewrite Zmod_small in H1 by omega. + congruence. +Qed. + +(** * Arithmetic and logical operations over machine integers *) + +Definition eq (x y: int) : bool := + if zeq (unsigned x) (unsigned y) then true else false. +Definition lt (x y: int) : bool := + if zlt (signed x) (signed y) then true else false. +Definition ltu (x y: int) : bool := + if zlt (unsigned x) (unsigned y) then true else false. + +Definition neg (x: int) : int := repr (- unsigned x). + +Definition add (x y: int) : int := + repr (unsigned x + unsigned y). +Definition sub (x y: int) : int := + repr (unsigned x - unsigned y). +Definition mul (x y: int) : int := + repr (unsigned x * unsigned y). + +Definition divs (x y: int) : int := + repr (Z.quot (signed x) (signed y)). +Definition mods (x y: int) : int := + repr (Z.rem (signed x) (signed y)). + +Definition divu (x y: int) : int := + repr (unsigned x / unsigned y). +Definition modu (x y: int) : int := + repr ((unsigned x) mod (unsigned y)). + +(** Bitwise boolean operations. *) + +Definition and (x y: int): int := repr (Z.land (unsigned x) (unsigned y)). +Definition or (x y: int): int := repr (Z.lor (unsigned x) (unsigned y)). +Definition xor (x y: int) : int := repr (Z.lxor (unsigned x) (unsigned y)). + +Definition not (x: int) : int := xor x mone. + +(** Shifts and rotates. *) + +Definition shl (x y: int): int := repr (Z.shiftl (unsigned x) (unsigned y)). +Definition shru (x y: int): int := repr (Z.shiftr (unsigned x) (unsigned y)). +Definition shr (x y: int): int := repr (Z.shiftr (signed x) (unsigned y)). + +Definition rol (x y: int) : int := + let n := (unsigned y) mod zwordsize in + repr (Z.lor (Z.shiftl (unsigned x) n) (Z.shiftr (unsigned x) (zwordsize - n))). +Definition ror (x y: int) : int := + let n := (unsigned y) mod zwordsize in + repr (Z.lor (Z.shiftr (unsigned x) n) (Z.shiftl (unsigned x) (zwordsize - n))). + +Definition rolm (x a m: int): int := and (rol x a) m. + +(** Viewed as signed divisions by powers of two, [shrx] rounds towards + zero, while [shr] rounds towards minus infinity. *) + +Definition shrx (x y: int): int := + divs x (shl one y). + +(** High half of full multiply. *) + +Definition mulhu (x y: int): int := repr ((unsigned x * unsigned y) / modulus). +Definition mulhs (x y: int): int := repr ((signed x * signed y) / modulus). + +(** Condition flags *) + +Definition negative (x: int): int := + if lt x zero then one else zero. + +Definition add_carry (x y cin: int): int := + if zlt (unsigned x + unsigned y + unsigned cin) modulus then zero else one. + +Definition add_overflow (x y cin: int): int := + let s := signed x + signed y + signed cin in + if zle min_signed s && zle s max_signed then zero else one. + +Definition sub_borrow (x y bin: int): int := + if zlt (unsigned x - unsigned y - unsigned bin) 0 then one else zero. + +Definition sub_overflow (x y bin: int): int := + let s := signed x - signed y - signed bin in + if zle min_signed s && zle s max_signed then zero else one. + +(** [shr_carry x y] is 1 if [x] is negative and at least one 1 bit is shifted away. *) + +Definition shr_carry (x y: int) : int := + if lt x zero && negb (eq (and x (sub (shl one y) one)) zero) + then one else zero. + +(** Zero and sign extensions *) + +Definition Zshiftin (b: bool) (x: Z) : Z := + if b then Z.succ_double x else Z.double x. + +(** In pseudo-code: +<< + Fixpoint Zzero_ext (n: Z) (x: Z) : Z := + if zle n 0 then + 0 + else + Zshiftin (Z.odd x) (Zzero_ext (Z.pred n) (Z.div2 x)). + Fixpoint Zsign_ext (n: Z) (x: Z) : Z := + if zle n 1 then + if Z.odd x then -1 else 0 + else + Zshiftin (Z.odd x) (Zzero_ext (Z.pred n) (Z.div2 x)). +>> + We encode this [nat]-like recursion using the [Z.iter] iteration + function, in order to make the [Zzero_ext] and [Zsign_ext] + functions efficiently executable within Coq. +*) + +Definition Zzero_ext (n: Z) (x: Z) : Z := + Z.iter n + (fun rec x => Zshiftin (Z.odd x) (rec (Z.div2 x))) + (fun x => 0) + x. + +Definition Zsign_ext (n: Z) (x: Z) : Z := + Z.iter (Z.pred n) + (fun rec x => Zshiftin (Z.odd x) (rec (Z.div2 x))) + (fun x => if Z.odd x then -1 else 0) + x. + +Definition zero_ext (n: Z) (x: int) : int := repr (Zzero_ext n (unsigned x)). + +Definition sign_ext (n: Z) (x: int) : int := repr (Zsign_ext n (unsigned x)). + +(** Decomposition of a number as a sum of powers of two. *) + +Fixpoint Z_one_bits (n: nat) (x: Z) (i: Z) {struct n}: list Z := + match n with + | O => nil + | S m => + if Z.odd x + then i :: Z_one_bits m (Z.div2 x) (i+1) + else Z_one_bits m (Z.div2 x) (i+1) + end. + +Definition one_bits (x: int) : list int := + List.map repr (Z_one_bits wordsize (unsigned x) 0). + +(** Recognition of powers of two. *) + +Definition is_power2 (x: int) : option int := + match Z_one_bits wordsize (unsigned x) 0 with + | i :: nil => Some (repr i) + | _ => None + end. + +(** Comparisons. *) + +Definition cmp (c: comparison) (x y: int) : bool := + match c with + | Ceq => eq x y + | Cne => negb (eq x y) + | Clt => lt x y + | Cle => negb (lt y x) + | Cgt => lt y x + | Cge => negb (lt x y) + end. + +Definition cmpu (c: comparison) (x y: int) : bool := + match c with + | Ceq => eq x y + | Cne => negb (eq x y) + | Clt => ltu x y + | Cle => negb (ltu y x) + | Cgt => ltu y x + | Cge => negb (ltu x y) + end. + +Definition is_false (x: int) : Prop := x = zero. +Definition is_true (x: int) : Prop := x <> zero. +Definition notbool (x: int) : int := if eq x zero then one else zero. + +(** x86-style extended division and modulus *) + +Definition divmodu2 (nhi nlo: int) (d: int) : option (int * int) := + if eq_dec d zero then None else + (let (q, r) := Z.div_eucl (unsigned nhi * modulus + unsigned nlo) (unsigned d) in + if zle q max_unsigned then Some(repr q, repr r) else None). + +Definition divmods2 (nhi nlo: int) (d: int) : option (int * int) := + if eq_dec d zero then None else + (let (q, r) := Z.quotrem (signed nhi * modulus + unsigned nlo) (signed d) in + if zle min_signed q && zle q max_signed then Some(repr q, repr r) else None). + +(** * Properties of integers and integer arithmetic *) + +(** ** Properties of [modulus], [max_unsigned], etc. *) + +Remark half_modulus_power: + half_modulus = two_p (zwordsize - 1). +Proof. + unfold half_modulus. rewrite modulus_power. + set (ws1 := zwordsize - 1). + replace (zwordsize) with (Zsucc ws1). + rewrite two_p_S. rewrite Zmult_comm. apply Z_div_mult. omega. + unfold ws1. generalize wordsize_pos; omega. + unfold ws1. omega. +Qed. + +Remark half_modulus_modulus: modulus = 2 * half_modulus. +Proof. + rewrite half_modulus_power. rewrite modulus_power. + rewrite <- two_p_S. apply f_equal. omega. + generalize wordsize_pos; omega. +Qed. + +(** Relative positions, from greatest to smallest: +<< + max_unsigned + max_signed + 2*wordsize-1 + wordsize + 0 + min_signed +>> +*) + +Remark half_modulus_pos: half_modulus > 0. +Proof. + rewrite half_modulus_power. apply two_p_gt_ZERO. generalize wordsize_pos; omega. +Qed. + +Remark min_signed_neg: min_signed < 0. +Proof. + unfold min_signed. generalize half_modulus_pos. omega. +Qed. + +Remark max_signed_pos: max_signed >= 0. +Proof. + unfold max_signed. generalize half_modulus_pos. omega. +Qed. + +Remark wordsize_max_unsigned: zwordsize <= max_unsigned. +Proof. + assert (zwordsize < modulus). + rewrite modulus_power. apply two_p_strict. + generalize wordsize_pos. omega. + unfold max_unsigned. omega. +Qed. + +Remark two_wordsize_max_unsigned: 2 * zwordsize - 1 <= max_unsigned. +Proof. + assert (2 * zwordsize - 1 < modulus). + rewrite modulus_power. apply two_p_strict_2. generalize wordsize_pos; omega. + unfold max_unsigned; omega. +Qed. + +Remark max_signed_unsigned: max_signed < max_unsigned. +Proof. + unfold max_signed, max_unsigned. rewrite half_modulus_modulus. + generalize half_modulus_pos. omega. +Qed. + +Lemma unsigned_repr_eq: + forall x, unsigned (repr x) = Zmod x modulus. +Proof. + intros. simpl. apply Z_mod_modulus_eq. +Qed. + +Lemma signed_repr_eq: + forall x, signed (repr x) = if zlt (Zmod x modulus) half_modulus then Zmod x modulus else Zmod x modulus - modulus. +Proof. + intros. unfold signed. rewrite unsigned_repr_eq. auto. +Qed. + +(** ** Modulo arithmetic *) + +(** We define and state properties of equality and arithmetic modulo a + positive integer. *) + +Section EQ_MODULO. + +Variable modul: Z. +Hypothesis modul_pos: modul > 0. + +Definition eqmod (x y: Z) : Prop := exists k, x = k * modul + y. + +Lemma eqmod_refl: forall x, eqmod x x. +Proof. + intros; red. exists 0. omega. +Qed. + +Lemma eqmod_refl2: forall x y, x = y -> eqmod x y. +Proof. + intros. subst y. apply eqmod_refl. +Qed. + +Lemma eqmod_sym: forall x y, eqmod x y -> eqmod y x. +Proof. + intros x y [k EQ]; red. exists (-k). subst x. ring. +Qed. + +Lemma eqmod_trans: forall x y z, eqmod x y -> eqmod y z -> eqmod x z. +Proof. + intros x y z [k1 EQ1] [k2 EQ2]; red. + exists (k1 + k2). subst x; subst y. ring. +Qed. + +Lemma eqmod_small_eq: + forall x y, eqmod x y -> 0 <= x < modul -> 0 <= y < modul -> x = y. +Proof. + intros x y [k EQ] I1 I2. + generalize (Zdiv_unique _ _ _ _ EQ I2). intro. + rewrite (Zdiv_small x modul I1) in H. subst k. omega. +Qed. + +Lemma eqmod_mod_eq: + forall x y, eqmod x y -> x mod modul = y mod modul. +Proof. + intros x y [k EQ]. subst x. + rewrite Zplus_comm. apply Z_mod_plus. auto. +Qed. + +Lemma eqmod_mod: + forall x, eqmod x (x mod modul). +Proof. + intros; red. exists (x / modul). + rewrite Zmult_comm. apply Z_div_mod_eq. auto. +Qed. + +Lemma eqmod_add: + forall a b c d, eqmod a b -> eqmod c d -> eqmod (a + c) (b + d). +Proof. + intros a b c d [k1 EQ1] [k2 EQ2]; red. + subst a; subst c. exists (k1 + k2). ring. +Qed. + +Lemma eqmod_neg: + forall x y, eqmod x y -> eqmod (-x) (-y). +Proof. + intros x y [k EQ]; red. exists (-k). rewrite EQ. ring. +Qed. + +Lemma eqmod_sub: + forall a b c d, eqmod a b -> eqmod c d -> eqmod (a - c) (b - d). +Proof. + intros a b c d [k1 EQ1] [k2 EQ2]; red. + subst a; subst c. exists (k1 - k2). ring. +Qed. + +Lemma eqmod_mult: + forall a b c d, eqmod a c -> eqmod b d -> eqmod (a * b) (c * d). +Proof. + intros a b c d [k1 EQ1] [k2 EQ2]; red. + subst a; subst b. + exists (k1 * k2 * modul + c * k2 + k1 * d). + ring. +Qed. + +End EQ_MODULO. + +Lemma eqmod_divides: + forall n m x y, eqmod n x y -> Zdivide m n -> eqmod m x y. +Proof. + intros. destruct H as [k1 EQ1]. destruct H0 as [k2 EQ2]. + exists (k1*k2). rewrite <- Zmult_assoc. rewrite <- EQ2. auto. +Qed. + +(** We then specialize these definitions to equality modulo + $2^{wordsize}$ #2wordsize#. *) + +Hint Resolve modulus_pos: ints. + +Definition eqm := eqmod modulus. + +Lemma eqm_refl: forall x, eqm x x. +Proof (eqmod_refl modulus). +Hint Resolve eqm_refl: ints. + +Lemma eqm_refl2: + forall x y, x = y -> eqm x y. +Proof (eqmod_refl2 modulus). +Hint Resolve eqm_refl2: ints. + +Lemma eqm_sym: forall x y, eqm x y -> eqm y x. +Proof (eqmod_sym modulus). +Hint Resolve eqm_sym: ints. + +Lemma eqm_trans: forall x y z, eqm x y -> eqm y z -> eqm x z. +Proof (eqmod_trans modulus). +Hint Resolve eqm_trans: ints. + +Lemma eqm_small_eq: + forall x y, eqm x y -> 0 <= x < modulus -> 0 <= y < modulus -> x = y. +Proof (eqmod_small_eq modulus). +Hint Resolve eqm_small_eq: ints. + +Lemma eqm_add: + forall a b c d, eqm a b -> eqm c d -> eqm (a + c) (b + d). +Proof (eqmod_add modulus). +Hint Resolve eqm_add: ints. + +Lemma eqm_neg: + forall x y, eqm x y -> eqm (-x) (-y). +Proof (eqmod_neg modulus). +Hint Resolve eqm_neg: ints. + +Lemma eqm_sub: + forall a b c d, eqm a b -> eqm c d -> eqm (a - c) (b - d). +Proof (eqmod_sub modulus). +Hint Resolve eqm_sub: ints. + +Lemma eqm_mult: + forall a b c d, eqm a c -> eqm b d -> eqm (a * b) (c * d). +Proof (eqmod_mult modulus). +Hint Resolve eqm_mult: ints. + +(** ** Properties of the coercions between [Z] and [int] *) + +Lemma eqm_samerepr: forall x y, eqm x y -> repr x = repr y. +Proof. + intros. unfold repr. apply mkint_eq. + rewrite !Z_mod_modulus_eq. apply eqmod_mod_eq. auto with ints. exact H. +Qed. + +Lemma eqm_unsigned_repr: + forall z, eqm z (unsigned (repr z)). +Proof. + unfold eqm; intros. rewrite unsigned_repr_eq. apply eqmod_mod. auto with ints. +Qed. +Hint Resolve eqm_unsigned_repr: ints. + +Lemma eqm_unsigned_repr_l: + forall a b, eqm a b -> eqm (unsigned (repr a)) b. +Proof. + intros. apply eqm_trans with a. + apply eqm_sym. apply eqm_unsigned_repr. auto. +Qed. +Hint Resolve eqm_unsigned_repr_l: ints. + +Lemma eqm_unsigned_repr_r: + forall a b, eqm a b -> eqm a (unsigned (repr b)). +Proof. + intros. apply eqm_trans with b. auto. + apply eqm_unsigned_repr. +Qed. +Hint Resolve eqm_unsigned_repr_r: ints. + +Lemma eqm_signed_unsigned: + forall x, eqm (signed x) (unsigned x). +Proof. + intros; red. unfold signed. set (y := unsigned x). + case (zlt y half_modulus); intro. + apply eqmod_refl. red; exists (-1); ring. +Qed. + +Theorem unsigned_range: + forall i, 0 <= unsigned i < modulus. +Proof. + destruct i. simpl. omega. +Qed. +Hint Resolve unsigned_range: ints. + +Theorem unsigned_range_2: + forall i, 0 <= unsigned i <= max_unsigned. +Proof. + intro; unfold max_unsigned. + generalize (unsigned_range i). omega. +Qed. +Hint Resolve unsigned_range_2: ints. + +Theorem signed_range: + forall i, min_signed <= signed i <= max_signed. +Proof. + intros. unfold signed. + generalize (unsigned_range i). set (n := unsigned i). intros. + case (zlt n half_modulus); intro. + unfold max_signed. generalize min_signed_neg. omega. + unfold min_signed, max_signed. + rewrite half_modulus_modulus in *. omega. +Qed. + +Theorem repr_unsigned: + forall i, repr (unsigned i) = i. +Proof. + destruct i; simpl. unfold repr. apply mkint_eq. + rewrite Z_mod_modulus_eq. apply Zmod_small; omega. +Qed. +Hint Resolve repr_unsigned: ints. + +Lemma repr_signed: + forall i, repr (signed i) = i. +Proof. + intros. transitivity (repr (unsigned i)). + apply eqm_samerepr. apply eqm_signed_unsigned. auto with ints. +Qed. +Hint Resolve repr_signed: ints. + +Opaque repr. + +Lemma eqm_repr_eq: forall x y, eqm x (unsigned y) -> repr x = y. +Proof. + intros. rewrite <- (repr_unsigned y). apply eqm_samerepr; auto. +Qed. + +Theorem unsigned_repr: + forall z, 0 <= z <= max_unsigned -> unsigned (repr z) = z. +Proof. + intros. rewrite unsigned_repr_eq. + apply Zmod_small. unfold max_unsigned in H. omega. +Qed. +Hint Resolve unsigned_repr: ints. + +Theorem signed_repr: + forall z, min_signed <= z <= max_signed -> signed (repr z) = z. +Proof. + intros. unfold signed. destruct (zle 0 z). + replace (unsigned (repr z)) with z. + rewrite zlt_true. auto. unfold max_signed in H. omega. + symmetry. apply unsigned_repr. generalize max_signed_unsigned. omega. + pose (z' := z + modulus). + replace (repr z) with (repr z'). + replace (unsigned (repr z')) with z'. + rewrite zlt_false. unfold z'. omega. + unfold z'. unfold min_signed in H. + rewrite half_modulus_modulus. omega. + symmetry. apply unsigned_repr. + unfold z', max_unsigned. unfold min_signed, max_signed in H. + rewrite half_modulus_modulus. omega. + apply eqm_samerepr. unfold z'; red. exists 1. omega. +Qed. + +Theorem signed_eq_unsigned: + forall x, unsigned x <= max_signed -> signed x = unsigned x. +Proof. + intros. unfold signed. destruct (zlt (unsigned x) half_modulus). + auto. unfold max_signed in H. omegaContradiction. +Qed. + +Theorem signed_positive: + forall x, signed x >= 0 <-> unsigned x <= max_signed. +Proof. + intros. unfold signed, max_signed. + generalize (unsigned_range x) half_modulus_modulus half_modulus_pos; intros. + destruct (zlt (unsigned x) half_modulus); omega. +Qed. + +(** ** Properties of zero, one, minus one *) + +Theorem unsigned_zero: unsigned zero = 0. +Proof. + unfold zero; rewrite unsigned_repr_eq. apply Zmod_0_l. +Qed. + +Theorem unsigned_one: unsigned one = 1. +Proof. + unfold one; rewrite unsigned_repr_eq. apply Zmod_small. split. omega. + unfold modulus. replace wordsize with (S(pred wordsize)). + rewrite two_power_nat_S. generalize (two_power_nat_pos (pred wordsize)). + omega. + generalize wordsize_pos. unfold zwordsize. omega. +Qed. + +Theorem unsigned_mone: unsigned mone = modulus - 1. +Proof. + unfold mone; rewrite unsigned_repr_eq. + replace (-1) with ((modulus - 1) + (-1) * modulus). + rewrite Z_mod_plus_full. apply Zmod_small. + generalize modulus_pos. omega. omega. +Qed. + +Theorem signed_zero: signed zero = 0. +Proof. + unfold signed. rewrite unsigned_zero. apply zlt_true. generalize half_modulus_pos; omega. +Qed. + +Theorem signed_one: zwordsize > 1 -> signed one = 1. +Proof. + intros. unfold signed. rewrite unsigned_one. apply zlt_true. + change 1 with (two_p 0). rewrite half_modulus_power. apply two_p_monotone_strict. omega. +Qed. + +Theorem signed_mone: signed mone = -1. +Proof. + unfold signed. rewrite unsigned_mone. + rewrite zlt_false. omega. + rewrite half_modulus_modulus. generalize half_modulus_pos. omega. +Qed. + +Theorem one_not_zero: one <> zero. +Proof. + assert (unsigned one <> unsigned zero). + rewrite unsigned_one; rewrite unsigned_zero; congruence. + congruence. +Qed. + +Theorem unsigned_repr_wordsize: + unsigned iwordsize = zwordsize. +Proof. + unfold iwordsize; rewrite unsigned_repr_eq. apply Zmod_small. + generalize wordsize_pos wordsize_max_unsigned; unfold max_unsigned; omega. +Qed. + +(** ** Properties of equality *) + +Theorem eq_sym: + forall x y, eq x y = eq y x. +Proof. + intros; unfold eq. case (zeq (unsigned x) (unsigned y)); intro. + rewrite e. rewrite zeq_true. auto. + rewrite zeq_false. auto. auto. +Qed. + +Theorem eq_spec: forall (x y: int), if eq x y then x = y else x <> y. +Proof. + intros; unfold eq. case (eq_dec x y); intro. + subst y. rewrite zeq_true. auto. + rewrite zeq_false. auto. + destruct x; destruct y. + simpl. red; intro. elim n. apply mkint_eq. auto. +Qed. + +Theorem eq_true: forall x, eq x x = true. +Proof. + intros. generalize (eq_spec x x); case (eq x x); intros; congruence. +Qed. + +Theorem eq_false: forall x y, x <> y -> eq x y = false. +Proof. + intros. generalize (eq_spec x y); case (eq x y); intros; congruence. +Qed. + +Theorem eq_signed: + forall x y, eq x y = if zeq (signed x) (signed y) then true else false. +Proof. + intros. predSpec eq eq_spec x y. + subst x. rewrite zeq_true; auto. + destruct (zeq (signed x) (signed y)); auto. + elim H. rewrite <- (repr_signed x). rewrite <- (repr_signed y). congruence. +Qed. + +(** ** Properties of addition *) + +Theorem add_unsigned: forall x y, add x y = repr (unsigned x + unsigned y). +Proof. intros; reflexivity. +Qed. + +Theorem add_signed: forall x y, add x y = repr (signed x + signed y). +Proof. + intros. rewrite add_unsigned. apply eqm_samerepr. + apply eqm_add; apply eqm_sym; apply eqm_signed_unsigned. +Qed. + +Theorem add_commut: forall x y, add x y = add y x. +Proof. intros; unfold add. decEq. omega. Qed. + +Theorem add_zero: forall x, add x zero = x. +Proof. + intros. unfold add. rewrite unsigned_zero. + rewrite Zplus_0_r. apply repr_unsigned. +Qed. + +Theorem add_zero_l: forall x, add zero x = x. +Proof. + intros. rewrite add_commut. apply add_zero. +Qed. + +Theorem add_assoc: forall x y z, add (add x y) z = add x (add y z). +Proof. + intros; unfold add. + set (x' := unsigned x). + set (y' := unsigned y). + set (z' := unsigned z). + apply eqm_samerepr. + apply eqm_trans with ((x' + y') + z'). + auto with ints. + rewrite <- Zplus_assoc. auto with ints. +Qed. + +Theorem add_permut: forall x y z, add x (add y z) = add y (add x z). +Proof. + intros. rewrite (add_commut y z). rewrite <- add_assoc. apply add_commut. +Qed. + +Theorem add_neg_zero: forall x, add x (neg x) = zero. +Proof. + intros; unfold add, neg, zero. apply eqm_samerepr. + replace 0 with (unsigned x + (- (unsigned x))). + auto with ints. omega. +Qed. + +Theorem unsigned_add_carry: + forall x y, + unsigned (add x y) = unsigned x + unsigned y - unsigned (add_carry x y zero) * modulus. +Proof. + intros. + unfold add, add_carry. rewrite unsigned_zero. rewrite Zplus_0_r. + rewrite unsigned_repr_eq. + generalize (unsigned_range x) (unsigned_range y). intros. + destruct (zlt (unsigned x + unsigned y) modulus). + rewrite unsigned_zero. apply Zmod_unique with 0. omega. omega. + rewrite unsigned_one. apply Zmod_unique with 1. omega. omega. +Qed. + +Corollary unsigned_add_either: + forall x y, + unsigned (add x y) = unsigned x + unsigned y + \/ unsigned (add x y) = unsigned x + unsigned y - modulus. +Proof. + intros. rewrite unsigned_add_carry. unfold add_carry. + rewrite unsigned_zero. rewrite Zplus_0_r. + destruct (zlt (unsigned x + unsigned y) modulus). + rewrite unsigned_zero. left; omega. + rewrite unsigned_one. right; omega. +Qed. + +(** ** Properties of negation *) + +Theorem neg_repr: forall z, neg (repr z) = repr (-z). +Proof. + intros; unfold neg. apply eqm_samerepr. auto with ints. +Qed. + +Theorem neg_zero: neg zero = zero. +Proof. + unfold neg. rewrite unsigned_zero. auto. +Qed. + +Theorem neg_involutive: forall x, neg (neg x) = x. +Proof. + intros; unfold neg. + apply eqm_repr_eq. eapply eqm_trans. apply eqm_neg. + apply eqm_unsigned_repr_l. apply eqm_refl. apply eqm_refl2. omega. +Qed. + +Theorem neg_add_distr: forall x y, neg(add x y) = add (neg x) (neg y). +Proof. + intros; unfold neg, add. apply eqm_samerepr. + apply eqm_trans with (- (unsigned x + unsigned y)). + auto with ints. + replace (- (unsigned x + unsigned y)) + with ((- unsigned x) + (- unsigned y)). + auto with ints. omega. +Qed. + +(** ** Properties of subtraction *) + +Theorem sub_zero_l: forall x, sub x zero = x. +Proof. + intros; unfold sub. rewrite unsigned_zero. + replace (unsigned x - 0) with (unsigned x) by omega. apply repr_unsigned. +Qed. + +Theorem sub_zero_r: forall x, sub zero x = neg x. +Proof. + intros; unfold sub, neg. rewrite unsigned_zero. auto. +Qed. + +Theorem sub_add_opp: forall x y, sub x y = add x (neg y). +Proof. + intros; unfold sub, add, neg. apply eqm_samerepr. + apply eqm_add; auto with ints. +Qed. + +Theorem sub_idem: forall x, sub x x = zero. +Proof. + intros; unfold sub. unfold zero. decEq. omega. +Qed. + +Theorem sub_add_l: forall x y z, sub (add x y) z = add (sub x z) y. +Proof. + intros. repeat rewrite sub_add_opp. + repeat rewrite add_assoc. decEq. apply add_commut. +Qed. + +Theorem sub_add_r: forall x y z, sub x (add y z) = add (sub x z) (neg y). +Proof. + intros. repeat rewrite sub_add_opp. + rewrite neg_add_distr. rewrite add_permut. apply add_commut. +Qed. + +Theorem sub_shifted: + forall x y z, + sub (add x z) (add y z) = sub x y. +Proof. + intros. rewrite sub_add_opp. rewrite neg_add_distr. + rewrite add_assoc. + rewrite (add_commut (neg y) (neg z)). + rewrite <- (add_assoc z). rewrite add_neg_zero. + rewrite (add_commut zero). rewrite add_zero. + symmetry. apply sub_add_opp. +Qed. + +Theorem sub_signed: + forall x y, sub x y = repr (signed x - signed y). +Proof. + intros. unfold sub. apply eqm_samerepr. + apply eqm_sub; apply eqm_sym; apply eqm_signed_unsigned. +Qed. + +Theorem unsigned_sub_borrow: + forall x y, + unsigned (sub x y) = unsigned x - unsigned y + unsigned (sub_borrow x y zero) * modulus. +Proof. + intros. + unfold sub, sub_borrow. rewrite unsigned_zero. rewrite Zminus_0_r. + rewrite unsigned_repr_eq. + generalize (unsigned_range x) (unsigned_range y). intros. + destruct (zlt (unsigned x - unsigned y) 0). + rewrite unsigned_one. apply Zmod_unique with (-1). omega. omega. + rewrite unsigned_zero. apply Zmod_unique with 0. omega. omega. +Qed. + +(** ** Properties of multiplication *) + +Theorem mul_commut: forall x y, mul x y = mul y x. +Proof. + intros; unfold mul. decEq. ring. +Qed. + +Theorem mul_zero: forall x, mul x zero = zero. +Proof. + intros; unfold mul. rewrite unsigned_zero. + unfold zero. decEq. ring. +Qed. + +Theorem mul_one: forall x, mul x one = x. +Proof. + intros; unfold mul. rewrite unsigned_one. + transitivity (repr (unsigned x)). decEq. ring. + apply repr_unsigned. +Qed. + +Theorem mul_mone: forall x, mul x mone = neg x. +Proof. + intros; unfold mul, neg. rewrite unsigned_mone. + apply eqm_samerepr. + replace (-unsigned x) with (0 - unsigned x) by omega. + replace (unsigned x * (modulus - 1)) with (unsigned x * modulus - unsigned x) by ring. + apply eqm_sub. exists (unsigned x). omega. apply eqm_refl. +Qed. + +Theorem mul_assoc: forall x y z, mul (mul x y) z = mul x (mul y z). +Proof. + intros; unfold mul. + set (x' := unsigned x). + set (y' := unsigned y). + set (z' := unsigned z). + apply eqm_samerepr. apply eqm_trans with ((x' * y') * z'). + auto with ints. + rewrite <- Zmult_assoc. auto with ints. +Qed. + +Theorem mul_add_distr_l: + forall x y z, mul (add x y) z = add (mul x z) (mul y z). +Proof. + intros; unfold mul, add. + apply eqm_samerepr. + set (x' := unsigned x). + set (y' := unsigned y). + set (z' := unsigned z). + apply eqm_trans with ((x' + y') * z'). + auto with ints. + replace ((x' + y') * z') with (x' * z' + y' * z'). + auto with ints. + ring. +Qed. + +Theorem mul_add_distr_r: + forall x y z, mul x (add y z) = add (mul x y) (mul x z). +Proof. + intros. rewrite mul_commut. rewrite mul_add_distr_l. + decEq; apply mul_commut. +Qed. + +Theorem neg_mul_distr_l: + forall x y, neg(mul x y) = mul (neg x) y. +Proof. + intros. unfold mul, neg. + set (x' := unsigned x). set (y' := unsigned y). + apply eqm_samerepr. apply eqm_trans with (- (x' * y')). + auto with ints. + replace (- (x' * y')) with ((-x') * y') by ring. + auto with ints. +Qed. + +Theorem neg_mul_distr_r: + forall x y, neg(mul x y) = mul x (neg y). +Proof. + intros. rewrite (mul_commut x y). rewrite (mul_commut x (neg y)). + apply neg_mul_distr_l. +Qed. + +Theorem mul_signed: + forall x y, mul x y = repr (signed x * signed y). +Proof. + intros; unfold mul. apply eqm_samerepr. + apply eqm_mult; apply eqm_sym; apply eqm_signed_unsigned. +Qed. + +(** ** Properties of division and modulus *) + +Lemma modu_divu_Euclid: + forall x y, y <> zero -> x = add (mul (divu x y) y) (modu x y). +Proof. + intros. unfold add, mul, divu, modu. + transitivity (repr (unsigned x)). auto with ints. + apply eqm_samerepr. + set (x' := unsigned x). set (y' := unsigned y). + apply eqm_trans with ((x' / y') * y' + x' mod y'). + apply eqm_refl2. rewrite Zmult_comm. apply Z_div_mod_eq. + generalize (unsigned_range y); intro. + assert (unsigned y <> 0). red; intro. + elim H. rewrite <- (repr_unsigned y). unfold zero. congruence. + unfold y'. omega. + auto with ints. +Qed. + +Theorem modu_divu: + forall x y, y <> zero -> modu x y = sub x (mul (divu x y) y). +Proof. + intros. + assert (forall a b c, a = add b c -> c = sub a b). + intros. subst a. rewrite sub_add_l. rewrite sub_idem. + rewrite add_commut. rewrite add_zero. auto. + apply H0. apply modu_divu_Euclid. auto. +Qed. + +Lemma mods_divs_Euclid: + forall x y, x = add (mul (divs x y) y) (mods x y). +Proof. + intros. unfold add, mul, divs, mods. + transitivity (repr (signed x)). auto with ints. + apply eqm_samerepr. + set (x' := signed x). set (y' := signed y). + apply eqm_trans with ((Z.quot x' y') * y' + Z.rem x' y'). + apply eqm_refl2. rewrite Zmult_comm. apply Z.quot_rem'. + apply eqm_add; auto with ints. + apply eqm_unsigned_repr_r. apply eqm_mult; auto with ints. + unfold y'. apply eqm_signed_unsigned. +Qed. + +Theorem mods_divs: + forall x y, mods x y = sub x (mul (divs x y) y). +Proof. + intros. + assert (forall a b c, a = add b c -> c = sub a b). + intros. subst a. rewrite sub_add_l. rewrite sub_idem. + rewrite add_commut. rewrite add_zero. auto. + apply H. apply mods_divs_Euclid. +Qed. + +Theorem divu_one: + forall x, divu x one = x. +Proof. + unfold divu; intros. rewrite unsigned_one. rewrite Zdiv_1_r. apply repr_unsigned. +Qed. + +Theorem modu_one: + forall x, modu x one = zero. +Proof. + intros. rewrite modu_divu. rewrite divu_one. rewrite mul_one. apply sub_idem. + apply one_not_zero. +Qed. + +Theorem divs_mone: + forall x, divs x mone = neg x. +Proof. + unfold divs, neg; intros. + rewrite signed_mone. + replace (Z.quot (signed x) (-1)) with (- (signed x)). + apply eqm_samerepr. apply eqm_neg. apply eqm_signed_unsigned. + set (x' := signed x). + set (one := 1). + change (-1) with (- one). rewrite Zquot_opp_r. + assert (Z.quot x' one = x'). + symmetry. apply Zquot_unique_full with 0. red. + change (Z.abs one) with 1. + destruct (zle 0 x'). left. omega. right. omega. + unfold one; ring. + congruence. +Qed. + +Theorem mods_mone: + forall x, mods x mone = zero. +Proof. + intros. rewrite mods_divs. rewrite divs_mone. + rewrite <- neg_mul_distr_l. rewrite mul_mone. rewrite neg_involutive. apply sub_idem. +Qed. + +Theorem divmodu2_divu_modu: + forall n d, + d <> zero -> divmodu2 zero n d = Some (divu n d, modu n d). +Proof. + unfold divmodu2, divu, modu; intros. + rewrite dec_eq_false by auto. + set (N := unsigned zero * modulus + unsigned n). + assert (E1: unsigned n = N) by (unfold N; rewrite unsigned_zero; ring). rewrite ! E1. + set (D := unsigned d). + set (Q := N / D); set (R := N mod D). + assert (E2: Z.div_eucl N D = (Q, R)). + { unfold Q, R, Z.div, Z.modulo. destruct (Z.div_eucl N D); auto. } + rewrite E2. rewrite zle_true. auto. + assert (unsigned d <> 0). + { red; intros. elim H. rewrite <- (repr_unsigned d). rewrite H0; auto. } + assert (0 < D). + { unfold D. generalize (unsigned_range d); intros. omega. } + assert (0 <= Q <= max_unsigned). + { unfold Q. apply Zdiv_interval_2. + rewrite <- E1; apply unsigned_range_2. + omega. unfold max_unsigned; generalize modulus_pos; omega. omega. } + omega. +Qed. + +Lemma unsigned_signed: + forall n, unsigned n = if lt n zero then signed n + modulus else signed n. +Proof. + intros. unfold lt. rewrite signed_zero. unfold signed. + generalize (unsigned_range n). rewrite half_modulus_modulus. intros. + destruct (zlt (unsigned n) half_modulus). +- rewrite zlt_false by omega. auto. +- rewrite zlt_true by omega. ring. +Qed. + +Theorem divmods2_divs_mods: + forall n d, + d <> zero -> n <> repr min_signed \/ d <> mone -> + divmods2 (if lt n zero then mone else zero) n d = Some (divs n d, mods n d). +Proof. + unfold divmods2, divs, mods; intros. + rewrite dec_eq_false by auto. + set (N := signed (if lt n zero then mone else zero) * modulus + unsigned n). + set (D := signed d). + assert (D <> 0). + { unfold D; red; intros. elim H. rewrite <- (repr_signed d). rewrite H1; auto. } + assert (N = signed n). + { unfold N. rewrite unsigned_signed. destruct (lt n zero). + rewrite signed_mone. ring. + rewrite signed_zero. ring. } + set (Q := Z.quot N D); set (R := Z.rem N D). + assert (E2: Z.quotrem N D = (Q, R)). + { unfold Q, R, Z.quot, Z.rem. destruct (Z.quotrem N D); auto. } + rewrite E2. + assert (min_signed <= N <= max_signed) by (rewrite H2; apply signed_range). + assert (min_signed <= Q <= max_signed). + { unfold Q. destruct (zeq D 1); [ | destruct (zeq D (-1))]. + - (* D = 1 *) + rewrite e. rewrite Z.quot_1_r; auto. + - (* D = -1 *) + rewrite e. change (-1) with (Z.opp 1). rewrite Z.quot_opp_r by omega. + rewrite Z.quot_1_r. + assert (N <> min_signed). + { red; intros; destruct H0. + + elim H0. rewrite <- (repr_signed n). rewrite <- H2. rewrite H4. auto. + + elim H0. rewrite <- (repr_signed d). unfold D in e; rewrite e; auto. } + unfold min_signed, max_signed in *. omega. + - (* |D| > 1 *) + assert (Z.abs (Z.quot N D) < half_modulus). + { rewrite <- Z.quot_abs by omega. apply Zquot_lt_upper_bound. + xomega. xomega. + apply Zle_lt_trans with (half_modulus * 1). + rewrite Z.mul_1_r. unfold min_signed, max_signed in H3; xomega. + apply Zmult_lt_compat_l. generalize half_modulus_pos; omega. xomega. } + rewrite Z.abs_lt in H4. + unfold min_signed, max_signed; omega. + } + unfold proj_sumbool; rewrite ! zle_true by omega; simpl. + unfold Q, R; rewrite H2; auto. +Qed. + +(** ** Bit-level properties *) + +(** ** Properties of bit-level operations over [Z] *) + +Remark Ztestbit_0: forall n, Z.testbit 0 n = false. +Proof Z.testbit_0_l. + +Remark Ztestbit_1: forall n, Z.testbit 1 n = zeq n 0. +Proof. + intros. destruct n; simpl; auto. +Qed. + +Remark Ztestbit_m1: forall n, 0 <= n -> Z.testbit (-1) n = true. +Proof. + intros. destruct n; simpl; auto. +Qed. + +Remark Zshiftin_spec: + forall b x, Zshiftin b x = 2 * x + (if b then 1 else 0). +Proof. + unfold Zshiftin; intros. destruct b. + - rewrite Z.succ_double_spec. omega. + - rewrite Z.double_spec. omega. +Qed. + +Remark Zshiftin_inj: + forall b1 x1 b2 x2, + Zshiftin b1 x1 = Zshiftin b2 x2 -> b1 = b2 /\ x1 = x2. +Proof. + intros. rewrite !Zshiftin_spec in H. + destruct b1; destruct b2. + split; [auto|omega]. + omegaContradiction. + omegaContradiction. + split; [auto|omega]. +Qed. + +Remark Zdecomp: + forall x, x = Zshiftin (Z.odd x) (Z.div2 x). +Proof. + intros. destruct x; simpl. + - auto. + - destruct p; auto. + - destruct p; auto. simpl. rewrite Pos.pred_double_succ. auto. +Qed. + +Remark Ztestbit_shiftin: + forall b x n, + 0 <= n -> + Z.testbit (Zshiftin b x) n = if zeq n 0 then b else Z.testbit x (Z.pred n). +Proof. + intros. rewrite Zshiftin_spec. destruct (zeq n 0). + - subst n. destruct b. + + apply Z.testbit_odd_0. + + rewrite Zplus_0_r. apply Z.testbit_even_0. + - assert (0 <= Z.pred n) by omega. + set (n' := Z.pred n) in *. + replace n with (Z.succ n') by (unfold n'; omega). + destruct b. + + apply Z.testbit_odd_succ; auto. + + rewrite Zplus_0_r. apply Z.testbit_even_succ; auto. +Qed. + +Remark Ztestbit_shiftin_base: + forall b x, Z.testbit (Zshiftin b x) 0 = b. +Proof. + intros. rewrite Ztestbit_shiftin. apply zeq_true. omega. +Qed. + +Remark Ztestbit_shiftin_succ: + forall b x n, 0 <= n -> Z.testbit (Zshiftin b x) (Z.succ n) = Z.testbit x n. +Proof. + intros. rewrite Ztestbit_shiftin. rewrite zeq_false. rewrite Z.pred_succ. auto. + omega. omega. +Qed. + +Remark Ztestbit_eq: + forall n x, 0 <= n -> + Z.testbit x n = if zeq n 0 then Z.odd x else Z.testbit (Z.div2 x) (Z.pred n). +Proof. + intros. rewrite (Zdecomp x) at 1. apply Ztestbit_shiftin; auto. +Qed. + +Remark Ztestbit_base: + forall x, Z.testbit x 0 = Z.odd x. +Proof. + intros. rewrite Ztestbit_eq. apply zeq_true. omega. +Qed. + +Remark Ztestbit_succ: + forall n x, 0 <= n -> Z.testbit x (Z.succ n) = Z.testbit (Z.div2 x) n. +Proof. + intros. rewrite Ztestbit_eq. rewrite zeq_false. rewrite Z.pred_succ. auto. + omega. omega. +Qed. + +Lemma eqmod_same_bits: + forall n x y, + (forall i, 0 <= i < Z.of_nat n -> Z.testbit x i = Z.testbit y i) -> + eqmod (two_power_nat n) x y. +Proof. + induction n; intros. + - change (two_power_nat 0) with 1. exists (x-y); ring. + - rewrite two_power_nat_S. + assert (eqmod (two_power_nat n) (Z.div2 x) (Z.div2 y)). + apply IHn. intros. rewrite <- !Ztestbit_succ. apply H. rewrite inj_S; omega. + omega. omega. + destruct H0 as [k EQ]. + exists k. rewrite (Zdecomp x). rewrite (Zdecomp y). + replace (Z.odd y) with (Z.odd x). + rewrite EQ. rewrite !Zshiftin_spec. ring. + exploit (H 0). rewrite inj_S; omega. + rewrite !Ztestbit_base. auto. +Qed. + +Lemma eqm_same_bits: + forall x y, + (forall i, 0 <= i < zwordsize -> Z.testbit x i = Z.testbit y i) -> + eqm x y. +Proof (eqmod_same_bits wordsize). + +Lemma same_bits_eqmod: + forall n x y i, + eqmod (two_power_nat n) x y -> 0 <= i < Z.of_nat n -> + Z.testbit x i = Z.testbit y i. +Proof. + induction n; intros. + - simpl in H0. omegaContradiction. + - rewrite inj_S in H0. rewrite two_power_nat_S in H. + rewrite !(Ztestbit_eq i); intuition. + destruct H as [k EQ]. + assert (EQ': Zshiftin (Z.odd x) (Z.div2 x) = + Zshiftin (Z.odd y) (k * two_power_nat n + Z.div2 y)). + { + rewrite (Zdecomp x) in EQ. rewrite (Zdecomp y) in EQ. + rewrite EQ. rewrite !Zshiftin_spec. ring. + } + exploit Zshiftin_inj; eauto. intros [A B]. + destruct (zeq i 0). + + auto. + + apply IHn. exists k; auto. omega. +Qed. + +Lemma same_bits_eqm: + forall x y i, + eqm x y -> + 0 <= i < zwordsize -> + Z.testbit x i = Z.testbit y i. +Proof (same_bits_eqmod wordsize). + +Remark two_power_nat_infinity: + forall x, 0 <= x -> exists n, x < two_power_nat n. +Proof. + intros x0 POS0; pattern x0; apply natlike_ind; auto. + exists O. compute; auto. + intros. destruct H0 as [n LT]. exists (S n). rewrite two_power_nat_S. + generalize (two_power_nat_pos n). omega. +Qed. + +Lemma equal_same_bits: + forall x y, + (forall i, 0 <= i -> Z.testbit x i = Z.testbit y i) -> + x = y. +Proof. + intros. + set (z := if zlt x y then y - x else x - y). + assert (0 <= z). + unfold z; destruct (zlt x y); omega. + exploit (two_power_nat_infinity z); auto. intros [n LT]. + assert (eqmod (two_power_nat n) x y). + apply eqmod_same_bits. intros. apply H. tauto. + assert (eqmod (two_power_nat n) z 0). + unfold z. destruct (zlt x y). + replace 0 with (y - y) by omega. apply eqmod_sub. apply eqmod_refl. auto. + replace 0 with (x - x) by omega. apply eqmod_sub. apply eqmod_refl. apply eqmod_sym; auto. + assert (z = 0). + apply eqmod_small_eq with (two_power_nat n). auto. omega. generalize (two_power_nat_pos n); omega. + unfold z in H3. destruct (zlt x y); omega. +Qed. + +Lemma Z_one_complement: + forall i, 0 <= i -> + forall x, Z.testbit (-x-1) i = negb (Z.testbit x i). +Proof. + intros i0 POS0. pattern i0. apply Zlt_0_ind; auto. + intros i IND POS x. + rewrite (Zdecomp x). set (y := Z.div2 x). + replace (- Zshiftin (Z.odd x) y - 1) + with (Zshiftin (negb (Z.odd x)) (- y - 1)). + rewrite !Ztestbit_shiftin; auto. + destruct (zeq i 0). auto. apply IND. omega. + rewrite !Zshiftin_spec. destruct (Z.odd x); simpl negb; ring. +Qed. + +Lemma Ztestbit_above: + forall n x i, + 0 <= x < two_power_nat n -> + i >= Z.of_nat n -> + Z.testbit x i = false. +Proof. + induction n; intros. + - change (two_power_nat 0) with 1 in H. + replace x with 0 by omega. + apply Z.testbit_0_l. + - rewrite inj_S in H0. rewrite Ztestbit_eq. rewrite zeq_false. + apply IHn. rewrite two_power_nat_S in H. rewrite (Zdecomp x) in H. + rewrite Zshiftin_spec in H. destruct (Z.odd x); omega. + omega. omega. omega. +Qed. + +Lemma Ztestbit_above_neg: + forall n x i, + -two_power_nat n <= x < 0 -> + i >= Z.of_nat n -> + Z.testbit x i = true. +Proof. + intros. set (y := -x-1). + assert (Z.testbit y i = false). + apply Ztestbit_above with n. + unfold y; omega. auto. + unfold y in H1. rewrite Z_one_complement in H1. + change true with (negb false). rewrite <- H1. rewrite negb_involutive; auto. + omega. +Qed. + +Lemma Zsign_bit: + forall n x, + 0 <= x < two_power_nat (S n) -> + Z.testbit x (Z_of_nat n) = if zlt x (two_power_nat n) then false else true. +Proof. + induction n; intros. + - change (two_power_nat 1) with 2 in H. + assert (x = 0 \/ x = 1) by omega. + destruct H0; subst x; reflexivity. + - rewrite inj_S. rewrite Ztestbit_eq. rewrite zeq_false. rewrite Z.pred_succ. + rewrite IHn. rewrite two_power_nat_S. + destruct (zlt (Z.div2 x) (two_power_nat n)); rewrite (Zdecomp x); rewrite Zshiftin_spec. + rewrite zlt_true. auto. destruct (Z.odd x); omega. + rewrite zlt_false. auto. destruct (Z.odd x); omega. + rewrite (Zdecomp x) in H; rewrite Zshiftin_spec in H. + rewrite two_power_nat_S in H. destruct (Z.odd x); omega. + omega. omega. +Qed. + +Lemma Zshiftin_ind: + forall (P: Z -> Prop), + P 0 -> + (forall b x, 0 <= x -> P x -> P (Zshiftin b x)) -> + forall x, 0 <= x -> P x. +Proof. + intros. destruct x. + - auto. + - induction p. + + change (P (Zshiftin true (Z.pos p))). auto. + + change (P (Zshiftin false (Z.pos p))). auto. + + change (P (Zshiftin true 0)). apply H0. omega. auto. + - compute in H1. intuition congruence. +Qed. + +Lemma Zshiftin_pos_ind: + forall (P: Z -> Prop), + P 1 -> + (forall b x, 0 < x -> P x -> P (Zshiftin b x)) -> + forall x, 0 < x -> P x. +Proof. + intros. destruct x; simpl in H1; try discriminate. + induction p. + + change (P (Zshiftin true (Z.pos p))). auto. + + change (P (Zshiftin false (Z.pos p))). auto. + + auto. +Qed. + +Lemma Ztestbit_le: + forall x y, + 0 <= y -> + (forall i, 0 <= i -> Z.testbit x i = true -> Z.testbit y i = true) -> + x <= y. +Proof. + intros x y0 POS0; revert x; pattern y0; apply Zshiftin_ind; auto; intros. + - replace x with 0. omega. apply equal_same_bits; intros. + rewrite Ztestbit_0. destruct (Z.testbit x i) as [] eqn:E; auto. + exploit H; eauto. rewrite Ztestbit_0. auto. + - assert (Z.div2 x0 <= x). + { apply H0. intros. exploit (H1 (Zsucc i)). + omega. rewrite Ztestbit_succ; auto. rewrite Ztestbit_shiftin_succ; auto. + } + rewrite (Zdecomp x0). rewrite !Zshiftin_spec. + destruct (Z.odd x0) as [] eqn:E1; destruct b as [] eqn:E2; try omega. + exploit (H1 0). omega. rewrite Ztestbit_base; auto. + rewrite Ztestbit_shiftin_base. congruence. +Qed. + +(** ** Bit-level reasoning over type [int] *) + +Definition testbit (x: int) (i: Z) : bool := Z.testbit (unsigned x) i. + +Lemma testbit_repr: + forall x i, + 0 <= i < zwordsize -> + testbit (repr x) i = Z.testbit x i. +Proof. + intros. unfold testbit. apply same_bits_eqm; auto with ints. +Qed. + +Lemma same_bits_eq: + forall x y, + (forall i, 0 <= i < zwordsize -> testbit x i = testbit y i) -> + x = y. +Proof. + intros. rewrite <- (repr_unsigned x). rewrite <- (repr_unsigned y). + apply eqm_samerepr. apply eqm_same_bits. auto. +Qed. + +Lemma bits_above: + forall x i, i >= zwordsize -> testbit x i = false. +Proof. + intros. apply Ztestbit_above with wordsize; auto. apply unsigned_range. +Qed. + +Lemma bits_zero: + forall i, testbit zero i = false. +Proof. + intros. unfold testbit. rewrite unsigned_zero. apply Ztestbit_0. +Qed. + +Remark bits_one: forall n, testbit one n = zeq n 0. +Proof. + unfold testbit; intros. rewrite unsigned_one. apply Ztestbit_1. +Qed. + +Lemma bits_mone: + forall i, 0 <= i < zwordsize -> testbit mone i = true. +Proof. + intros. unfold mone. rewrite testbit_repr; auto. apply Ztestbit_m1. omega. +Qed. + +Hint Rewrite bits_zero bits_mone : ints. + +Ltac bit_solve := + intros; apply same_bits_eq; intros; autorewrite with ints; auto with bool. + +Lemma sign_bit_of_unsigned: + forall x, testbit x (zwordsize - 1) = if zlt (unsigned x) half_modulus then false else true. +Proof. + intros. unfold testbit. + set (ws1 := pred wordsize). + assert (zwordsize - 1 = Z_of_nat ws1). + unfold zwordsize, ws1, wordsize. + destruct WS.wordsize as [] eqn:E. + elim WS.wordsize_not_zero; auto. + rewrite inj_S. simpl. omega. + assert (half_modulus = two_power_nat ws1). + rewrite two_power_nat_two_p. rewrite <- H. apply half_modulus_power. + rewrite H; rewrite H0. + apply Zsign_bit. rewrite two_power_nat_S. rewrite <- H0. + rewrite <- half_modulus_modulus. apply unsigned_range. +Qed. + +Lemma bits_signed: + forall x i, 0 <= i -> + Z.testbit (signed x) i = testbit x (if zlt i zwordsize then i else zwordsize - 1). +Proof. + intros. + destruct (zlt i zwordsize). + - apply same_bits_eqm. apply eqm_signed_unsigned. omega. + - unfold signed. rewrite sign_bit_of_unsigned. destruct (zlt (unsigned x) half_modulus). + + apply Ztestbit_above with wordsize. apply unsigned_range. auto. + + apply Ztestbit_above_neg with wordsize. + fold modulus. generalize (unsigned_range x). omega. auto. +Qed. + +Lemma bits_le: + forall x y, + (forall i, 0 <= i < zwordsize -> testbit x i = true -> testbit y i = true) -> + unsigned x <= unsigned y. +Proof. + intros. apply Ztestbit_le. generalize (unsigned_range y); omega. + intros. fold (testbit y i). destruct (zlt i zwordsize). + apply H. omega. auto. + fold (testbit x i) in H1. rewrite bits_above in H1; auto. congruence. +Qed. + +(** ** Properties of bitwise and, or, xor *) + +Lemma bits_and: + forall x y i, 0 <= i < zwordsize -> + testbit (and x y) i = testbit x i && testbit y i. +Proof. + intros. unfold and. rewrite testbit_repr; auto. rewrite Z.land_spec; intuition. +Qed. + +Lemma bits_or: + forall x y i, 0 <= i < zwordsize -> + testbit (or x y) i = testbit x i || testbit y i. +Proof. + intros. unfold or. rewrite testbit_repr; auto. rewrite Z.lor_spec; intuition. +Qed. + +Lemma bits_xor: + forall x y i, 0 <= i < zwordsize -> + testbit (xor x y) i = xorb (testbit x i) (testbit y i). +Proof. + intros. unfold xor. rewrite testbit_repr; auto. rewrite Z.lxor_spec; intuition. +Qed. + +Lemma bits_not: + forall x i, 0 <= i < zwordsize -> + testbit (not x) i = negb (testbit x i). +Proof. + intros. unfold not. rewrite bits_xor; auto. rewrite bits_mone; auto. +Qed. + +Hint Rewrite bits_and bits_or bits_xor bits_not: ints. + +Theorem and_commut: forall x y, and x y = and y x. +Proof. + bit_solve. +Qed. + +Theorem and_assoc: forall x y z, and (and x y) z = and x (and y z). +Proof. + bit_solve. +Qed. + +Theorem and_zero: forall x, and x zero = zero. +Proof. + bit_solve. apply andb_b_false. +Qed. + +Corollary and_zero_l: forall x, and zero x = zero. +Proof. + intros. rewrite and_commut. apply and_zero. +Qed. + +Theorem and_mone: forall x, and x mone = x. +Proof. + bit_solve. apply andb_b_true. +Qed. + +Corollary and_mone_l: forall x, and mone x = x. +Proof. + intros. rewrite and_commut. apply and_mone. +Qed. + +Theorem and_idem: forall x, and x x = x. +Proof. + bit_solve. destruct (testbit x i); auto. +Qed. + +Theorem or_commut: forall x y, or x y = or y x. +Proof. + bit_solve. +Qed. + +Theorem or_assoc: forall x y z, or (or x y) z = or x (or y z). +Proof. + bit_solve. +Qed. + +Theorem or_zero: forall x, or x zero = x. +Proof. + bit_solve. +Qed. + +Corollary or_zero_l: forall x, or zero x = x. +Proof. + intros. rewrite or_commut. apply or_zero. +Qed. + +Theorem or_mone: forall x, or x mone = mone. +Proof. + bit_solve. +Qed. + +Theorem or_idem: forall x, or x x = x. +Proof. + bit_solve. destruct (testbit x i); auto. +Qed. + +Theorem and_or_distrib: + forall x y z, + and x (or y z) = or (and x y) (and x z). +Proof. + bit_solve. apply demorgan1. +Qed. + +Corollary and_or_distrib_l: + forall x y z, + and (or x y) z = or (and x z) (and y z). +Proof. + intros. rewrite (and_commut (or x y)). rewrite and_or_distrib. f_equal; apply and_commut. +Qed. + +Theorem or_and_distrib: + forall x y z, + or x (and y z) = and (or x y) (or x z). +Proof. + bit_solve. apply orb_andb_distrib_r. +Qed. + +Corollary or_and_distrib_l: + forall x y z, + or (and x y) z = and (or x z) (or y z). +Proof. + intros. rewrite (or_commut (and x y)). rewrite or_and_distrib. f_equal; apply or_commut. +Qed. + +Theorem and_or_absorb: forall x y, and x (or x y) = x. +Proof. + bit_solve. + assert (forall a b, a && (a || b) = a) by destr_bool. + auto. +Qed. + +Theorem or_and_absorb: forall x y, or x (and x y) = x. +Proof. + bit_solve. + assert (forall a b, a || (a && b) = a) by destr_bool. + auto. +Qed. + +Theorem xor_commut: forall x y, xor x y = xor y x. +Proof. + bit_solve. apply xorb_comm. +Qed. + +Theorem xor_assoc: forall x y z, xor (xor x y) z = xor x (xor y z). +Proof. + bit_solve. apply xorb_assoc. +Qed. + +Theorem xor_zero: forall x, xor x zero = x. +Proof. + bit_solve. apply xorb_false. +Qed. + +Corollary xor_zero_l: forall x, xor zero x = x. +Proof. + intros. rewrite xor_commut. apply xor_zero. +Qed. + +Theorem xor_idem: forall x, xor x x = zero. +Proof. + bit_solve. apply xorb_nilpotent. +Qed. + +Theorem xor_zero_one: xor zero one = one. +Proof. rewrite xor_commut. apply xor_zero. Qed. + +Theorem xor_one_one: xor one one = zero. +Proof. apply xor_idem. Qed. + +Theorem xor_zero_equal: forall x y, xor x y = zero -> x = y. +Proof. + intros. apply same_bits_eq; intros. + assert (xorb (testbit x i) (testbit y i) = false). + rewrite <- bits_xor; auto. rewrite H. apply bits_zero. + destruct (testbit x i); destruct (testbit y i); reflexivity || discriminate. +Qed. + +Theorem xor_is_zero: forall x y, eq (xor x y) zero = eq x y. +Proof. + intros. predSpec eq eq_spec (xor x y) zero. +- apply xor_zero_equal in H. subst y. rewrite eq_true; auto. +- predSpec eq eq_spec x y. ++ elim H; subst y; apply xor_idem. ++ auto. +Qed. + +Theorem and_xor_distrib: + forall x y z, + and x (xor y z) = xor (and x y) (and x z). +Proof. + bit_solve. + assert (forall a b c, a && (xorb b c) = xorb (a && b) (a && c)) by destr_bool. + auto. +Qed. + +Theorem and_le: + forall x y, unsigned (and x y) <= unsigned x. +Proof. + intros. apply bits_le; intros. + rewrite bits_and in H0; auto. rewrite andb_true_iff in H0. tauto. +Qed. + +Theorem or_le: + forall x y, unsigned x <= unsigned (or x y). +Proof. + intros. apply bits_le; intros. + rewrite bits_or; auto. rewrite H0; auto. +Qed. + +(** Properties of bitwise complement.*) + +Theorem not_involutive: + forall (x: int), not (not x) = x. +Proof. + intros. unfold not. rewrite xor_assoc. rewrite xor_idem. apply xor_zero. +Qed. + +Theorem not_zero: + not zero = mone. +Proof. + unfold not. rewrite xor_commut. apply xor_zero. +Qed. + +Theorem not_mone: + not mone = zero. +Proof. + rewrite <- (not_involutive zero). symmetry. decEq. apply not_zero. +Qed. + +Theorem not_or_and_not: + forall x y, not (or x y) = and (not x) (not y). +Proof. + bit_solve. apply negb_orb. +Qed. + +Theorem not_and_or_not: + forall x y, not (and x y) = or (not x) (not y). +Proof. + bit_solve. apply negb_andb. +Qed. + +Theorem and_not_self: + forall x, and x (not x) = zero. +Proof. + bit_solve. +Qed. + +Theorem or_not_self: + forall x, or x (not x) = mone. +Proof. + bit_solve. +Qed. + +Theorem xor_not_self: + forall x, xor x (not x) = mone. +Proof. + bit_solve. destruct (testbit x i); auto. +Qed. + +Lemma unsigned_not: + forall x, unsigned (not x) = max_unsigned - unsigned x. +Proof. + intros. transitivity (unsigned (repr(-unsigned x - 1))). + f_equal. bit_solve. rewrite testbit_repr; auto. symmetry. apply Z_one_complement. omega. + rewrite unsigned_repr_eq. apply Zmod_unique with (-1). + unfold max_unsigned. omega. + generalize (unsigned_range x). unfold max_unsigned. omega. +Qed. + +Theorem not_neg: + forall x, not x = add (neg x) mone. +Proof. + bit_solve. + rewrite <- (repr_unsigned x) at 1. unfold add. + rewrite !testbit_repr; auto. + transitivity (Z.testbit (-unsigned x - 1) i). + symmetry. apply Z_one_complement. omega. + apply same_bits_eqm; auto. + replace (-unsigned x - 1) with (-unsigned x + (-1)) by omega. + apply eqm_add. + unfold neg. apply eqm_unsigned_repr. + rewrite unsigned_mone. exists (-1). ring. +Qed. + +Theorem neg_not: + forall x, neg x = add (not x) one. +Proof. + intros. rewrite not_neg. rewrite add_assoc. + replace (add mone one) with zero. rewrite add_zero. auto. + apply eqm_samerepr. rewrite unsigned_mone. rewrite unsigned_one. + exists (-1). ring. +Qed. + +Theorem sub_add_not: + forall x y, sub x y = add (add x (not y)) one. +Proof. + intros. rewrite sub_add_opp. rewrite neg_not. + rewrite ! add_assoc. auto. +Qed. + +Theorem sub_add_not_3: + forall x y b, + b = zero \/ b = one -> + sub (sub x y) b = add (add x (not y)) (xor b one). +Proof. + intros. rewrite ! sub_add_not. rewrite ! add_assoc. f_equal. f_equal. + rewrite <- neg_not. rewrite <- sub_add_opp. destruct H; subst b. + rewrite xor_zero_l. rewrite sub_zero_l. auto. + rewrite xor_idem. rewrite sub_idem. auto. +Qed. + +Theorem sub_borrow_add_carry: + forall x y b, + b = zero \/ b = one -> + sub_borrow x y b = xor (add_carry x (not y) (xor b one)) one. +Proof. + intros. unfold sub_borrow, add_carry. rewrite unsigned_not. + replace (unsigned (xor b one)) with (1 - unsigned b). + destruct (zlt (unsigned x - unsigned y - unsigned b)). + rewrite zlt_true. rewrite xor_zero_l; auto. + unfold max_unsigned; omega. + rewrite zlt_false. rewrite xor_idem; auto. + unfold max_unsigned; omega. + destruct H; subst b. + rewrite xor_zero_l. rewrite unsigned_one, unsigned_zero; auto. + rewrite xor_idem. rewrite unsigned_one, unsigned_zero; auto. +Qed. + +(** Connections between [add] and bitwise logical operations. *) + +Lemma Z_add_is_or: + forall i, 0 <= i -> + forall x y, + (forall j, 0 <= j <= i -> Z.testbit x j && Z.testbit y j = false) -> + Z.testbit (x + y) i = Z.testbit x i || Z.testbit y i. +Proof. + intros i0 POS0. pattern i0. apply Zlt_0_ind; auto. + intros i IND POS x y EXCL. + rewrite (Zdecomp x) in *. rewrite (Zdecomp y) in *. + transitivity (Z.testbit (Zshiftin (Z.odd x || Z.odd y) (Z.div2 x + Z.div2 y)) i). + - f_equal. rewrite !Zshiftin_spec. + exploit (EXCL 0). omega. rewrite !Ztestbit_shiftin_base. intros. +Opaque Z.mul. + destruct (Z.odd x); destruct (Z.odd y); simpl in *; discriminate || ring. + - rewrite !Ztestbit_shiftin; auto. + destruct (zeq i 0). + + auto. + + apply IND. omega. intros. + exploit (EXCL (Z.succ j)). omega. + rewrite !Ztestbit_shiftin_succ. auto. + omega. omega. +Qed. + +Theorem add_is_or: + forall x y, + and x y = zero -> + add x y = or x y. +Proof. + bit_solve. unfold add. rewrite testbit_repr; auto. + apply Z_add_is_or. omega. + intros. + assert (testbit (and x y) j = testbit zero j) by congruence. + autorewrite with ints in H2. assumption. omega. +Qed. + +Theorem xor_is_or: + forall x y, and x y = zero -> xor x y = or x y. +Proof. + bit_solve. + assert (testbit (and x y) i = testbit zero i) by congruence. + autorewrite with ints in H1; auto. + destruct (testbit x i); destruct (testbit y i); simpl in *; congruence. +Qed. + +Theorem add_is_xor: + forall x y, + and x y = zero -> + add x y = xor x y. +Proof. + intros. rewrite xor_is_or; auto. apply add_is_or; auto. +Qed. + +Theorem add_and: + forall x y z, + and y z = zero -> + add (and x y) (and x z) = and x (or y z). +Proof. + intros. rewrite add_is_or. + rewrite and_or_distrib; auto. + rewrite (and_commut x y). + rewrite and_assoc. + repeat rewrite <- (and_assoc x). + rewrite (and_commut (and x x)). + rewrite <- and_assoc. + rewrite H. rewrite and_commut. apply and_zero. +Qed. + +(** ** Properties of shifts *) + +Lemma bits_shl: + forall x y i, + 0 <= i < zwordsize -> + testbit (shl x y) i = + if zlt i (unsigned y) then false else testbit x (i - unsigned y). +Proof. + intros. unfold shl. rewrite testbit_repr; auto. + destruct (zlt i (unsigned y)). + apply Z.shiftl_spec_low. auto. + apply Z.shiftl_spec_high. omega. omega. +Qed. + +Lemma bits_shru: + forall x y i, + 0 <= i < zwordsize -> + testbit (shru x y) i = + if zlt (i + unsigned y) zwordsize then testbit x (i + unsigned y) else false. +Proof. + intros. unfold shru. rewrite testbit_repr; auto. + rewrite Z.shiftr_spec. fold (testbit x (i + unsigned y)). + destruct (zlt (i + unsigned y) zwordsize). + auto. + apply bits_above; auto. + omega. +Qed. + +Lemma bits_shr: + forall x y i, + 0 <= i < zwordsize -> + testbit (shr x y) i = + testbit x (if zlt (i + unsigned y) zwordsize then i + unsigned y else zwordsize - 1). +Proof. + intros. unfold shr. rewrite testbit_repr; auto. + rewrite Z.shiftr_spec. apply bits_signed. + generalize (unsigned_range y); omega. + omega. +Qed. + +Hint Rewrite bits_shl bits_shru bits_shr: ints. + +Theorem shl_zero: forall x, shl x zero = x. +Proof. + bit_solve. rewrite unsigned_zero. rewrite zlt_false. f_equal; omega. omega. +Qed. + +Lemma bitwise_binop_shl: + forall f f' x y n, + (forall x y i, 0 <= i < zwordsize -> testbit (f x y) i = f' (testbit x i) (testbit y i)) -> + f' false false = false -> + f (shl x n) (shl y n) = shl (f x y) n. +Proof. + intros. apply same_bits_eq; intros. + rewrite H; auto. rewrite !bits_shl; auto. + destruct (zlt i (unsigned n)); auto. + rewrite H; auto. generalize (unsigned_range n); omega. +Qed. + +Theorem and_shl: + forall x y n, + and (shl x n) (shl y n) = shl (and x y) n. +Proof. + intros. apply bitwise_binop_shl with andb. exact bits_and. auto. +Qed. + +Theorem or_shl: + forall x y n, + or (shl x n) (shl y n) = shl (or x y) n. +Proof. + intros. apply bitwise_binop_shl with orb. exact bits_or. auto. +Qed. + +Theorem xor_shl: + forall x y n, + xor (shl x n) (shl y n) = shl (xor x y) n. +Proof. + intros. apply bitwise_binop_shl with xorb. exact bits_xor. auto. +Qed. + +Lemma ltu_inv: + forall x y, ltu x y = true -> 0 <= unsigned x < unsigned y. +Proof. + unfold ltu; intros. destruct (zlt (unsigned x) (unsigned y)). + split; auto. generalize (unsigned_range x); omega. + discriminate. +Qed. + +Lemma ltu_iwordsize_inv: + forall x, ltu x iwordsize = true -> 0 <= unsigned x < zwordsize. +Proof. + intros. generalize (ltu_inv _ _ H). rewrite unsigned_repr_wordsize. auto. +Qed. + +Theorem shl_shl: + forall x y z, + ltu y iwordsize = true -> + ltu z iwordsize = true -> + ltu (add y z) iwordsize = true -> + shl (shl x y) z = shl x (add y z). +Proof. + intros. + generalize (ltu_iwordsize_inv _ H) (ltu_iwordsize_inv _ H0); intros. + assert (unsigned (add y z) = unsigned y + unsigned z). + unfold add. apply unsigned_repr. + generalize two_wordsize_max_unsigned; omega. + apply same_bits_eq; intros. + rewrite bits_shl; auto. + destruct (zlt i (unsigned z)). + - rewrite bits_shl; auto. rewrite zlt_true. auto. omega. + - rewrite bits_shl. destruct (zlt (i - unsigned z) (unsigned y)). + + rewrite bits_shl; auto. rewrite zlt_true. auto. omega. + + rewrite bits_shl; auto. rewrite zlt_false. f_equal. omega. omega. + + omega. +Qed. + +Theorem sub_ltu: + forall x y, + ltu x y = true -> + 0 <= unsigned y - unsigned x <= unsigned y. +Proof. + intros. + generalize (ltu_inv x y H). intros . + split. omega. omega. +Qed. + +Theorem shru_zero: forall x, shru x zero = x. +Proof. + bit_solve. rewrite unsigned_zero. rewrite zlt_true. f_equal; omega. omega. +Qed. + +Lemma bitwise_binop_shru: + forall f f' x y n, + (forall x y i, 0 <= i < zwordsize -> testbit (f x y) i = f' (testbit x i) (testbit y i)) -> + f' false false = false -> + f (shru x n) (shru y n) = shru (f x y) n. +Proof. + intros. apply same_bits_eq; intros. + rewrite H; auto. rewrite !bits_shru; auto. + destruct (zlt (i + unsigned n) zwordsize); auto. + rewrite H; auto. generalize (unsigned_range n); omega. +Qed. + +Theorem and_shru: + forall x y n, + and (shru x n) (shru y n) = shru (and x y) n. +Proof. + intros. apply bitwise_binop_shru with andb; auto. exact bits_and. +Qed. + +Theorem or_shru: + forall x y n, + or (shru x n) (shru y n) = shru (or x y) n. +Proof. + intros. apply bitwise_binop_shru with orb; auto. exact bits_or. +Qed. + +Theorem xor_shru: + forall x y n, + xor (shru x n) (shru y n) = shru (xor x y) n. +Proof. + intros. apply bitwise_binop_shru with xorb; auto. exact bits_xor. +Qed. + +Theorem shru_shru: + forall x y z, + ltu y iwordsize = true -> + ltu z iwordsize = true -> + ltu (add y z) iwordsize = true -> + shru (shru x y) z = shru x (add y z). +Proof. + intros. + generalize (ltu_iwordsize_inv _ H) (ltu_iwordsize_inv _ H0); intros. + assert (unsigned (add y z) = unsigned y + unsigned z). + unfold add. apply unsigned_repr. + generalize two_wordsize_max_unsigned; omega. + apply same_bits_eq; intros. + rewrite bits_shru; auto. + destruct (zlt (i + unsigned z) zwordsize). + - rewrite bits_shru. destruct (zlt (i + unsigned z + unsigned y) zwordsize). + + rewrite bits_shru; auto. rewrite zlt_true. f_equal. omega. omega. + + rewrite bits_shru; auto. rewrite zlt_false. auto. omega. + + omega. + - rewrite bits_shru; auto. rewrite zlt_false. auto. omega. +Qed. + +Theorem shr_zero: forall x, shr x zero = x. +Proof. + bit_solve. rewrite unsigned_zero. rewrite zlt_true. f_equal; omega. omega. +Qed. + +Lemma bitwise_binop_shr: + forall f f' x y n, + (forall x y i, 0 <= i < zwordsize -> testbit (f x y) i = f' (testbit x i) (testbit y i)) -> + f (shr x n) (shr y n) = shr (f x y) n. +Proof. + intros. apply same_bits_eq; intros. + rewrite H; auto. rewrite !bits_shr; auto. + rewrite H; auto. + destruct (zlt (i + unsigned n) zwordsize). + generalize (unsigned_range n); omega. + omega. +Qed. + +Theorem and_shr: + forall x y n, + and (shr x n) (shr y n) = shr (and x y) n. +Proof. + intros. apply bitwise_binop_shr with andb. exact bits_and. +Qed. + +Theorem or_shr: + forall x y n, + or (shr x n) (shr y n) = shr (or x y) n. +Proof. + intros. apply bitwise_binop_shr with orb. exact bits_or. +Qed. + +Theorem xor_shr: + forall x y n, + xor (shr x n) (shr y n) = shr (xor x y) n. +Proof. + intros. apply bitwise_binop_shr with xorb. exact bits_xor. +Qed. + +Theorem shr_shr: + forall x y z, + ltu y iwordsize = true -> + ltu z iwordsize = true -> + ltu (add y z) iwordsize = true -> + shr (shr x y) z = shr x (add y z). +Proof. + intros. + generalize (ltu_iwordsize_inv _ H) (ltu_iwordsize_inv _ H0); intros. + assert (unsigned (add y z) = unsigned y + unsigned z). + unfold add. apply unsigned_repr. + generalize two_wordsize_max_unsigned; omega. + apply same_bits_eq; intros. + rewrite !bits_shr; auto. f_equal. + destruct (zlt (i + unsigned z) zwordsize). + rewrite H4. replace (i + (unsigned y + unsigned z)) with (i + unsigned z + unsigned y) by omega. auto. + rewrite (zlt_false _ (i + unsigned (add y z))). + destruct (zlt (zwordsize - 1 + unsigned y) zwordsize); omega. + omega. + destruct (zlt (i + unsigned z) zwordsize); omega. +Qed. + +Theorem and_shr_shru: + forall x y z, + and (shr x z) (shru y z) = shru (and x y) z. +Proof. + intros. apply same_bits_eq; intros. + rewrite bits_and; auto. rewrite bits_shr; auto. rewrite !bits_shru; auto. + destruct (zlt (i + unsigned z) zwordsize). + - rewrite bits_and; auto. generalize (unsigned_range z); omega. + - apply andb_false_r. +Qed. + +Theorem shr_and_shru_and: + forall x y z, + shru (shl z y) y = z -> + and (shr x y) z = and (shru x y) z. +Proof. + intros. + rewrite <- H. + rewrite and_shru. rewrite and_shr_shru. auto. +Qed. + +Theorem shru_lt_zero: + forall x, + shru x (repr (zwordsize - 1)) = if lt x zero then one else zero. +Proof. + intros. apply same_bits_eq; intros. + rewrite bits_shru; auto. + rewrite unsigned_repr. + destruct (zeq i 0). + subst i. rewrite Zplus_0_l. rewrite zlt_true. + rewrite sign_bit_of_unsigned. + unfold lt. rewrite signed_zero. unfold signed. + destruct (zlt (unsigned x) half_modulus). + rewrite zlt_false. auto. generalize (unsigned_range x); omega. + rewrite zlt_true. unfold one; rewrite testbit_repr; auto. + generalize (unsigned_range x); omega. + omega. + rewrite zlt_false. + unfold testbit. rewrite Ztestbit_eq. rewrite zeq_false. + destruct (lt x zero). + rewrite unsigned_one. simpl Z.div2. rewrite Z.testbit_0_l; auto. + rewrite unsigned_zero. simpl Z.div2. rewrite Z.testbit_0_l; auto. + auto. omega. omega. + generalize wordsize_max_unsigned; omega. +Qed. + +Theorem shr_lt_zero: + forall x, + shr x (repr (zwordsize - 1)) = if lt x zero then mone else zero. +Proof. + intros. apply same_bits_eq; intros. + rewrite bits_shr; auto. + rewrite unsigned_repr. + transitivity (testbit x (zwordsize - 1)). + f_equal. destruct (zlt (i + (zwordsize - 1)) zwordsize); omega. + rewrite sign_bit_of_unsigned. + unfold lt. rewrite signed_zero. unfold signed. + destruct (zlt (unsigned x) half_modulus). + rewrite zlt_false. rewrite bits_zero; auto. generalize (unsigned_range x); omega. + rewrite zlt_true. rewrite bits_mone; auto. generalize (unsigned_range x); omega. + generalize wordsize_max_unsigned; omega. +Qed. + +(** ** Properties of rotations *) + +Lemma bits_rol: + forall x y i, + 0 <= i < zwordsize -> + testbit (rol x y) i = testbit x ((i - unsigned y) mod zwordsize). +Proof. + intros. unfold rol. + exploit (Z_div_mod_eq (unsigned y) zwordsize). apply wordsize_pos. + set (j := unsigned y mod zwordsize). set (k := unsigned y / zwordsize). + intros EQ. + exploit (Z_mod_lt (unsigned y) zwordsize). apply wordsize_pos. + fold j. intros RANGE. + rewrite testbit_repr; auto. + rewrite Z.lor_spec. rewrite Z.shiftr_spec. 2: omega. + destruct (zlt i j). + - rewrite Z.shiftl_spec_low; auto. simpl. + unfold testbit. f_equal. + symmetry. apply Zmod_unique with (-k - 1). + rewrite EQ. ring. + omega. + - rewrite Z.shiftl_spec_high. + fold (testbit x (i + (zwordsize - j))). + rewrite bits_above. rewrite orb_false_r. + fold (testbit x (i - j)). + f_equal. symmetry. apply Zmod_unique with (-k). + rewrite EQ. ring. + omega. omega. omega. omega. +Qed. + +Lemma bits_ror: + forall x y i, + 0 <= i < zwordsize -> + testbit (ror x y) i = testbit x ((i + unsigned y) mod zwordsize). +Proof. + intros. unfold ror. + exploit (Z_div_mod_eq (unsigned y) zwordsize). apply wordsize_pos. + set (j := unsigned y mod zwordsize). set (k := unsigned y / zwordsize). + intros EQ. + exploit (Z_mod_lt (unsigned y) zwordsize). apply wordsize_pos. + fold j. intros RANGE. + rewrite testbit_repr; auto. + rewrite Z.lor_spec. rewrite Z.shiftr_spec. 2: omega. + destruct (zlt (i + j) zwordsize). + - rewrite Z.shiftl_spec_low; auto. rewrite orb_false_r. + unfold testbit. f_equal. + symmetry. apply Zmod_unique with k. + rewrite EQ. ring. + omega. omega. + - rewrite Z.shiftl_spec_high. + fold (testbit x (i + j)). + rewrite bits_above. simpl. + unfold testbit. f_equal. + symmetry. apply Zmod_unique with (k + 1). + rewrite EQ. ring. + omega. omega. omega. omega. +Qed. + +Hint Rewrite bits_rol bits_ror: ints. + +Theorem shl_rolm: + forall x n, + ltu n iwordsize = true -> + shl x n = rolm x n (shl mone n). +Proof. + intros. generalize (ltu_inv _ _ H). rewrite unsigned_repr_wordsize; intros. + unfold rolm. apply same_bits_eq; intros. + rewrite bits_and; auto. rewrite !bits_shl; auto. rewrite bits_rol; auto. + destruct (zlt i (unsigned n)). + - rewrite andb_false_r; auto. + - generalize (unsigned_range n); intros. + rewrite bits_mone. rewrite andb_true_r. f_equal. + symmetry. apply Zmod_small. omega. + omega. +Qed. + +Theorem shru_rolm: + forall x n, + ltu n iwordsize = true -> + shru x n = rolm x (sub iwordsize n) (shru mone n). +Proof. + intros. generalize (ltu_inv _ _ H). rewrite unsigned_repr_wordsize; intros. + unfold rolm. apply same_bits_eq; intros. + rewrite bits_and; auto. rewrite !bits_shru; auto. rewrite bits_rol; auto. + destruct (zlt (i + unsigned n) zwordsize). + - generalize (unsigned_range n); intros. + rewrite bits_mone. rewrite andb_true_r. f_equal. + unfold sub. rewrite unsigned_repr. rewrite unsigned_repr_wordsize. + symmetry. apply Zmod_unique with (-1). ring. omega. + rewrite unsigned_repr_wordsize. generalize wordsize_max_unsigned. omega. + omega. + - rewrite andb_false_r; auto. +Qed. + +Theorem rol_zero: + forall x, + rol x zero = x. +Proof. + bit_solve. f_equal. rewrite unsigned_zero. rewrite Zminus_0_r. + apply Zmod_small; auto. +Qed. + +Lemma bitwise_binop_rol: + forall f f' x y n, + (forall x y i, 0 <= i < zwordsize -> testbit (f x y) i = f' (testbit x i) (testbit y i)) -> + rol (f x y) n = f (rol x n) (rol y n). +Proof. + intros. apply same_bits_eq; intros. + rewrite H; auto. rewrite !bits_rol; auto. rewrite H; auto. + apply Z_mod_lt. apply wordsize_pos. +Qed. + +Theorem rol_and: + forall x y n, + rol (and x y) n = and (rol x n) (rol y n). +Proof. + intros. apply bitwise_binop_rol with andb. exact bits_and. +Qed. + +Theorem rol_or: + forall x y n, + rol (or x y) n = or (rol x n) (rol y n). +Proof. + intros. apply bitwise_binop_rol with orb. exact bits_or. +Qed. + +Theorem rol_xor: + forall x y n, + rol (xor x y) n = xor (rol x n) (rol y n). +Proof. + intros. apply bitwise_binop_rol with xorb. exact bits_xor. +Qed. + +Theorem rol_rol: + forall x n m, + Zdivide zwordsize modulus -> + rol (rol x n) m = rol x (modu (add n m) iwordsize). +Proof. + bit_solve. f_equal. apply eqmod_mod_eq. apply wordsize_pos. + set (M := unsigned m); set (N := unsigned n). + apply eqmod_trans with (i - M - N). + apply eqmod_sub. + apply eqmod_sym. apply eqmod_mod. apply wordsize_pos. + apply eqmod_refl. + replace (i - M - N) with (i - (M + N)) by omega. + apply eqmod_sub. + apply eqmod_refl. + apply eqmod_trans with (Zmod (unsigned n + unsigned m) zwordsize). + replace (M + N) with (N + M) by omega. apply eqmod_mod. apply wordsize_pos. + unfold modu, add. fold M; fold N. rewrite unsigned_repr_wordsize. + assert (forall a, eqmod zwordsize a (unsigned (repr a))). + intros. eapply eqmod_divides. apply eqm_unsigned_repr. assumption. + eapply eqmod_trans. 2: apply H1. + apply eqmod_refl2. apply eqmod_mod_eq. apply wordsize_pos. auto. + apply Z_mod_lt. apply wordsize_pos. +Qed. + +Theorem rolm_zero: + forall x m, + rolm x zero m = and x m. +Proof. + intros. unfold rolm. rewrite rol_zero. auto. +Qed. + +Theorem rolm_rolm: + forall x n1 m1 n2 m2, + Zdivide zwordsize modulus -> + rolm (rolm x n1 m1) n2 m2 = + rolm x (modu (add n1 n2) iwordsize) + (and (rol m1 n2) m2). +Proof. + intros. + unfold rolm. rewrite rol_and. rewrite and_assoc. + rewrite rol_rol. reflexivity. auto. +Qed. + +Theorem or_rolm: + forall x n m1 m2, + or (rolm x n m1) (rolm x n m2) = rolm x n (or m1 m2). +Proof. + intros; unfold rolm. symmetry. apply and_or_distrib. +Qed. + +Theorem ror_rol: + forall x y, + ltu y iwordsize = true -> + ror x y = rol x (sub iwordsize y). +Proof. + intros. + generalize (ltu_iwordsize_inv _ H); intros. + apply same_bits_eq; intros. + rewrite bits_ror; auto. rewrite bits_rol; auto. f_equal. + unfold sub. rewrite unsigned_repr. rewrite unsigned_repr_wordsize. + apply eqmod_mod_eq. apply wordsize_pos. exists 1. ring. + rewrite unsigned_repr_wordsize. + generalize wordsize_pos; generalize wordsize_max_unsigned; omega. +Qed. + +Theorem ror_rol_neg: + forall x y, (zwordsize | modulus) -> ror x y = rol x (neg y). +Proof. + intros. apply same_bits_eq; intros. + rewrite bits_ror by auto. rewrite bits_rol by auto. + f_equal. apply eqmod_mod_eq. omega. + apply eqmod_trans with (i - (- unsigned y)). + apply eqmod_refl2; omega. + apply eqmod_sub. apply eqmod_refl. + apply eqmod_divides with modulus. + apply eqm_unsigned_repr. auto. +Qed. + +Theorem or_ror: + forall x y z, + ltu y iwordsize = true -> + ltu z iwordsize = true -> + add y z = iwordsize -> + ror x z = or (shl x y) (shru x z). +Proof. + intros. + generalize (ltu_iwordsize_inv _ H) (ltu_iwordsize_inv _ H0); intros. + unfold ror, or, shl, shru. apply same_bits_eq; intros. + rewrite !testbit_repr; auto. + rewrite !Z.lor_spec. rewrite orb_comm. f_equal; apply same_bits_eqm; auto. + - apply eqm_unsigned_repr_r. apply eqm_refl2. f_equal. + rewrite Zmod_small; auto. + assert (unsigned (add y z) = zwordsize). + rewrite H1. apply unsigned_repr_wordsize. + unfold add in H5. rewrite unsigned_repr in H5. + omega. + generalize two_wordsize_max_unsigned; omega. + - apply eqm_unsigned_repr_r. apply eqm_refl2. f_equal. + apply Zmod_small; auto. +Qed. + +(** ** Properties of [Z_one_bits] and [is_power2]. *) + +Fixpoint powerserie (l: list Z): Z := + match l with + | nil => 0 + | x :: xs => two_p x + powerserie xs + end. + +Lemma Z_one_bits_powerserie: + forall x, 0 <= x < modulus -> x = powerserie (Z_one_bits wordsize x 0). +Proof. + assert (forall n x i, + 0 <= i -> + 0 <= x < two_power_nat n -> + x * two_p i = powerserie (Z_one_bits n x i)). + { + induction n; intros. + simpl. rewrite two_power_nat_O in H0. + assert (x = 0) by omega. subst x. omega. + rewrite two_power_nat_S in H0. simpl Z_one_bits. + rewrite (Zdecomp x) in H0. rewrite Zshiftin_spec in H0. + assert (EQ: Z.div2 x * two_p (i + 1) = powerserie (Z_one_bits n (Z.div2 x) (i + 1))). + apply IHn. omega. + destruct (Z.odd x); omega. + rewrite two_p_is_exp in EQ. change (two_p 1) with 2 in EQ. + rewrite (Zdecomp x) at 1. rewrite Zshiftin_spec. + destruct (Z.odd x); simpl powerserie; rewrite <- EQ; ring. + omega. omega. + } + intros. rewrite <- H. change (two_p 0) with 1. omega. + omega. exact H0. +Qed. + +Lemma Z_one_bits_range: + forall x i, In i (Z_one_bits wordsize x 0) -> 0 <= i < zwordsize. +Proof. + assert (forall n x i j, + In j (Z_one_bits n x i) -> i <= j < i + Z_of_nat n). + { + induction n; simpl In. + tauto. + intros x i j. rewrite inj_S. + assert (In j (Z_one_bits n (Z.div2 x) (i + 1)) -> i <= j < i + Z.succ (Z.of_nat n)). + intros. exploit IHn; eauto. omega. + destruct (Z.odd x); simpl. + intros [A|B]. subst j. omega. auto. + auto. + } + intros. generalize (H wordsize x 0 i H0). fold zwordsize; omega. +Qed. + +Lemma is_power2_rng: + forall n logn, + is_power2 n = Some logn -> + 0 <= unsigned logn < zwordsize. +Proof. + intros n logn. unfold is_power2. + generalize (Z_one_bits_range (unsigned n)). + destruct (Z_one_bits wordsize (unsigned n) 0). + intros; discriminate. + destruct l. + intros. injection H0; intro; subst logn; clear H0. + assert (0 <= z < zwordsize). + apply H. auto with coqlib. + rewrite unsigned_repr. auto. generalize wordsize_max_unsigned; omega. + intros; discriminate. +Qed. + +Theorem is_power2_range: + forall n logn, + is_power2 n = Some logn -> ltu logn iwordsize = true. +Proof. + intros. unfold ltu. rewrite unsigned_repr_wordsize. + apply zlt_true. generalize (is_power2_rng _ _ H). tauto. +Qed. + +Lemma is_power2_correct: + forall n logn, + is_power2 n = Some logn -> + unsigned n = two_p (unsigned logn). +Proof. + intros n logn. unfold is_power2. + generalize (Z_one_bits_powerserie (unsigned n) (unsigned_range n)). + generalize (Z_one_bits_range (unsigned n)). + destruct (Z_one_bits wordsize (unsigned n) 0). + intros; discriminate. + destruct l. + intros. simpl in H0. injection H1; intros; subst logn; clear H1. + rewrite unsigned_repr. replace (two_p z) with (two_p z + 0). + auto. omega. elim (H z); intros. + generalize wordsize_max_unsigned; omega. + auto with coqlib. + intros; discriminate. +Qed. + +Remark two_p_range: + forall n, + 0 <= n < zwordsize -> + 0 <= two_p n <= max_unsigned. +Proof. + intros. split. + assert (two_p n > 0). apply two_p_gt_ZERO. omega. omega. + generalize (two_p_monotone_strict _ _ H). + unfold zwordsize; rewrite <- two_power_nat_two_p. + unfold max_unsigned, modulus. omega. +Qed. + +Remark Z_one_bits_zero: + forall n i, Z_one_bits n 0 i = nil. +Proof. + induction n; intros; simpl; auto. +Qed. + +Remark Z_one_bits_two_p: + forall n x i, + 0 <= x < Z_of_nat n -> + Z_one_bits n (two_p x) i = (i + x) :: nil. +Proof. + induction n; intros; simpl. simpl in H. omegaContradiction. + rewrite inj_S in H. + assert (x = 0 \/ 0 < x) by omega. destruct H0. + subst x; simpl. decEq. omega. apply Z_one_bits_zero. + assert (Z.odd (two_p x) = false /\ Z.div2 (two_p x) = two_p (x-1)). + apply Zshiftin_inj. rewrite <- Zdecomp. rewrite !Zshiftin_spec. + rewrite <- two_p_S. rewrite Zplus_0_r. f_equal; omega. omega. + destruct H1 as [A B]; rewrite A; rewrite B. + rewrite IHn. f_equal; omega. omega. +Qed. + +Lemma is_power2_two_p: + forall n, 0 <= n < zwordsize -> + is_power2 (repr (two_p n)) = Some (repr n). +Proof. + intros. unfold is_power2. rewrite unsigned_repr. + rewrite Z_one_bits_two_p. auto. auto. + apply two_p_range. auto. +Qed. + +(** ** Relation between bitwise operations and multiplications / divisions by powers of 2 *) + +(** Left shifts and multiplications by powers of 2. *) + +Lemma Zshiftl_mul_two_p: + forall x n, 0 <= n -> Z.shiftl x n = x * two_p n. +Proof. + intros. destruct n; simpl. + - omega. + - pattern p. apply Pos.peano_ind. + + change (two_power_pos 1) with 2. simpl. ring. + + intros. rewrite Pos.iter_succ. rewrite H0. + rewrite Pplus_one_succ_l. rewrite two_power_pos_is_exp. + change (two_power_pos 1) with 2. ring. + - compute in H. congruence. +Qed. + +Lemma shl_mul_two_p: + forall x y, + shl x y = mul x (repr (two_p (unsigned y))). +Proof. + intros. unfold shl, mul. apply eqm_samerepr. + rewrite Zshiftl_mul_two_p. auto with ints. + generalize (unsigned_range y); omega. +Qed. + +Theorem shl_mul: + forall x y, + shl x y = mul x (shl one y). +Proof. + intros. + assert (shl one y = repr (two_p (unsigned y))). + { + rewrite shl_mul_two_p. rewrite mul_commut. rewrite mul_one. auto. + } + rewrite H. apply shl_mul_two_p. +Qed. + +Theorem mul_pow2: + forall x n logn, + is_power2 n = Some logn -> + mul x n = shl x logn. +Proof. + intros. generalize (is_power2_correct n logn H); intro. + rewrite shl_mul_two_p. rewrite <- H0. rewrite repr_unsigned. + auto. +Qed. + +Theorem shifted_or_is_add: + forall x y n, + 0 <= n < zwordsize -> + unsigned y < two_p n -> + or (shl x (repr n)) y = repr(unsigned x * two_p n + unsigned y). +Proof. + intros. rewrite <- add_is_or. + - unfold add. apply eqm_samerepr. apply eqm_add; auto with ints. + rewrite shl_mul_two_p. unfold mul. apply eqm_unsigned_repr_l. + apply eqm_mult; auto with ints. apply eqm_unsigned_repr_l. + apply eqm_refl2. rewrite unsigned_repr. auto. + generalize wordsize_max_unsigned; omega. + - bit_solve. + rewrite unsigned_repr. + destruct (zlt i n). + + auto. + + replace (testbit y i) with false. apply andb_false_r. + symmetry. unfold testbit. + assert (EQ: Z.of_nat (Z.to_nat n) = n) by (apply Z2Nat.id; omega). + apply Ztestbit_above with (Z.to_nat n). + rewrite <- EQ in H0. rewrite <- two_power_nat_two_p in H0. + generalize (unsigned_range y); omega. + rewrite EQ; auto. + + generalize wordsize_max_unsigned; omega. +Qed. + +(** Unsigned right shifts and unsigned divisions by powers of 2. *) + +Lemma Zshiftr_div_two_p: + forall x n, 0 <= n -> Z.shiftr x n = x / two_p n. +Proof. + intros. destruct n; unfold Z.shiftr; simpl. + - rewrite Zdiv_1_r. auto. + - pattern p. apply Pos.peano_ind. + + change (two_power_pos 1) with 2. simpl. apply Zdiv2_div. + + intros. rewrite Pos.iter_succ. rewrite H0. + rewrite Pplus_one_succ_l. rewrite two_power_pos_is_exp. + change (two_power_pos 1) with 2. + rewrite Zdiv2_div. rewrite Zmult_comm. apply Zdiv_Zdiv. + rewrite two_power_pos_nat. apply two_power_nat_pos. omega. + - compute in H. congruence. +Qed. + +Lemma shru_div_two_p: + forall x y, + shru x y = repr (unsigned x / two_p (unsigned y)). +Proof. + intros. unfold shru. + rewrite Zshiftr_div_two_p. auto. + generalize (unsigned_range y); omega. +Qed. + +Theorem divu_pow2: + forall x n logn, + is_power2 n = Some logn -> + divu x n = shru x logn. +Proof. + intros. generalize (is_power2_correct n logn H). intro. + symmetry. unfold divu. rewrite H0. apply shru_div_two_p. +Qed. + +(** Signed right shifts and signed divisions by powers of 2. *) + +Lemma shr_div_two_p: + forall x y, + shr x y = repr (signed x / two_p (unsigned y)). +Proof. + intros. unfold shr. + rewrite Zshiftr_div_two_p. auto. + generalize (unsigned_range y); omega. +Qed. + +Theorem divs_pow2: + forall x n logn, + is_power2 n = Some logn -> + divs x n = shrx x logn. +Proof. + intros. generalize (is_power2_correct _ _ H); intro. + unfold shrx. rewrite shl_mul_two_p. + rewrite mul_commut. rewrite mul_one. + rewrite <- H0. rewrite repr_unsigned. auto. +Qed. + +(** Unsigned modulus over [2^n] is masking with [2^n-1]. *) + +Lemma Ztestbit_mod_two_p: + forall n x i, + 0 <= n -> 0 <= i -> + Z.testbit (x mod (two_p n)) i = if zlt i n then Z.testbit x i else false. +Proof. + intros n0 x i N0POS. revert x i; pattern n0; apply natlike_ind; auto. + - intros. change (two_p 0) with 1. rewrite Zmod_1_r. rewrite Z.testbit_0_l. + rewrite zlt_false; auto. omega. + - intros. rewrite two_p_S; auto. + replace (x0 mod (2 * two_p x)) + with (Zshiftin (Z.odd x0) (Z.div2 x0 mod two_p x)). + rewrite Ztestbit_shiftin; auto. rewrite (Ztestbit_eq i x0); auto. destruct (zeq i 0). + + rewrite zlt_true; auto. omega. + + rewrite H0. destruct (zlt (Z.pred i) x). + * rewrite zlt_true; auto. omega. + * rewrite zlt_false; auto. omega. + * omega. + + rewrite (Zdecomp x0) at 3. set (x1 := Z.div2 x0). symmetry. + apply Zmod_unique with (x1 / two_p x). + rewrite !Zshiftin_spec. rewrite Zplus_assoc. f_equal. + transitivity (2 * (two_p x * (x1 / two_p x) + x1 mod two_p x)). + f_equal. apply Z_div_mod_eq. apply two_p_gt_ZERO; auto. + ring. + rewrite Zshiftin_spec. exploit (Z_mod_lt x1 (two_p x)). apply two_p_gt_ZERO; auto. + destruct (Z.odd x0); omega. +Qed. + +Corollary Ztestbit_two_p_m1: + forall n i, 0 <= n -> 0 <= i -> + Z.testbit (two_p n - 1) i = if zlt i n then true else false. +Proof. + intros. replace (two_p n - 1) with ((-1) mod (two_p n)). + rewrite Ztestbit_mod_two_p; auto. destruct (zlt i n); auto. apply Ztestbit_m1; auto. + apply Zmod_unique with (-1). ring. + exploit (two_p_gt_ZERO n). auto. omega. +Qed. + +Theorem modu_and: + forall x n logn, + is_power2 n = Some logn -> + modu x n = and x (sub n one). +Proof. + intros. generalize (is_power2_correct _ _ H); intro. + generalize (is_power2_rng _ _ H); intro. + apply same_bits_eq; intros. + rewrite bits_and; auto. + unfold sub. rewrite testbit_repr; auto. + rewrite H0. rewrite unsigned_one. + unfold modu. rewrite testbit_repr; auto. rewrite H0. + rewrite Ztestbit_mod_two_p. rewrite Ztestbit_two_p_m1. + destruct (zlt i (unsigned logn)). + rewrite andb_true_r; auto. + rewrite andb_false_r; auto. + tauto. tauto. tauto. tauto. +Qed. + +(** ** Properties of [shrx] (signed division by a power of 2) *) + +Lemma Zquot_Zdiv: + forall x y, + y > 0 -> + Z.quot x y = if zlt x 0 then (x + y - 1) / y else x / y. +Proof. + intros. destruct (zlt x 0). + - symmetry. apply Zquot_unique_full with ((x + y - 1) mod y - (y - 1)). + + red. right; split. omega. + exploit (Z_mod_lt (x + y - 1) y); auto. + rewrite Z.abs_eq. omega. omega. + + transitivity ((y * ((x + y - 1) / y) + (x + y - 1) mod y) - (y-1)). + rewrite <- Z_div_mod_eq. ring. auto. ring. + - apply Zquot_Zdiv_pos; omega. +Qed. + +Theorem shrx_zero: + forall x, zwordsize > 1 -> shrx x zero = x. +Proof. + intros. unfold shrx. rewrite shl_zero. unfold divs. rewrite signed_one by auto. + rewrite Z.quot_1_r. apply repr_signed. +Qed. + +Theorem shrx_shr: + forall x y, + ltu y (repr (zwordsize - 1)) = true -> + shrx x y = shr (if lt x zero then add x (sub (shl one y) one) else x) y. +Proof. + intros. + set (uy := unsigned y). + assert (0 <= uy < zwordsize - 1). + generalize (ltu_inv _ _ H). rewrite unsigned_repr. auto. + generalize wordsize_pos wordsize_max_unsigned; omega. + rewrite shr_div_two_p. unfold shrx. unfold divs. + assert (shl one y = repr (two_p uy)). + transitivity (mul one (repr (two_p uy))). + symmetry. apply mul_pow2. replace y with (repr uy). + apply is_power2_two_p. omega. apply repr_unsigned. + rewrite mul_commut. apply mul_one. + assert (two_p uy > 0). apply two_p_gt_ZERO. omega. + assert (two_p uy < half_modulus). + rewrite half_modulus_power. + apply two_p_monotone_strict. auto. + assert (two_p uy < modulus). + rewrite modulus_power. apply two_p_monotone_strict. omega. + assert (unsigned (shl one y) = two_p uy). + rewrite H1. apply unsigned_repr. unfold max_unsigned. omega. + assert (signed (shl one y) = two_p uy). + rewrite H1. apply signed_repr. + unfold max_signed. generalize min_signed_neg. omega. + rewrite H6. + rewrite Zquot_Zdiv; auto. + unfold lt. rewrite signed_zero. + destruct (zlt (signed x) 0); auto. + rewrite add_signed. + assert (signed (sub (shl one y) one) = two_p uy - 1). + unfold sub. rewrite H5. rewrite unsigned_one. + apply signed_repr. + generalize min_signed_neg. unfold max_signed. omega. + rewrite H7. rewrite signed_repr. f_equal. f_equal. omega. + generalize (signed_range x). intros. + assert (two_p uy - 1 <= max_signed). unfold max_signed. omega. omega. +Qed. + +Theorem shrx_shr_2: + forall x y, + ltu y (repr (zwordsize - 1)) = true -> + shrx x y = shr (add x (shru (shr x (repr (zwordsize - 1))) (sub iwordsize y))) y. +Proof. + intros. + rewrite shrx_shr by auto. f_equal. + rewrite shr_lt_zero. destruct (lt x zero). +- set (uy := unsigned y). + generalize (unsigned_range y); fold uy; intros. + assert (0 <= uy < zwordsize - 1). + generalize (ltu_inv _ _ H). rewrite unsigned_repr. auto. + generalize wordsize_pos wordsize_max_unsigned; omega. + assert (two_p uy < modulus). + rewrite modulus_power. apply two_p_monotone_strict. omega. + f_equal. rewrite shl_mul_two_p. fold uy. rewrite mul_commut. rewrite mul_one. + unfold sub. rewrite unsigned_one. rewrite unsigned_repr. + rewrite unsigned_repr_wordsize. fold uy. + apply same_bits_eq; intros. rewrite bits_shru by auto. + rewrite testbit_repr by auto. rewrite Ztestbit_two_p_m1 by omega. + rewrite unsigned_repr by (generalize wordsize_max_unsigned; omega). + destruct (zlt i uy). + rewrite zlt_true by omega. rewrite bits_mone by omega. auto. + rewrite zlt_false by omega. auto. + assert (two_p uy > 0) by (apply two_p_gt_ZERO; omega). unfold max_unsigned; omega. +- replace (shru zero (sub iwordsize y)) with zero. + rewrite add_zero; auto. + bit_solve. destruct (zlt (i + unsigned (sub iwordsize y)) zwordsize); auto. +Qed. + +Lemma Zdiv_shift: + forall x y, y > 0 -> + (x + (y - 1)) / y = x / y + if zeq (Zmod x y) 0 then 0 else 1. +Proof. + intros. generalize (Z_div_mod_eq x y H). generalize (Z_mod_lt x y H). + set (q := x / y). set (r := x mod y). intros. + destruct (zeq r 0). + apply Zdiv_unique with (y - 1). rewrite H1. rewrite e. ring. omega. + apply Zdiv_unique with (r - 1). rewrite H1. ring. omega. +Qed. + +Theorem shrx_carry: + forall x y, + ltu y (repr (zwordsize - 1)) = true -> + shrx x y = add (shr x y) (shr_carry x y). +Proof. + intros. rewrite shrx_shr; auto. unfold shr_carry. + unfold lt. set (sx := signed x). rewrite signed_zero. + destruct (zlt sx 0); simpl. + 2: rewrite add_zero; auto. + set (uy := unsigned y). + assert (0 <= uy < zwordsize - 1). + generalize (ltu_inv _ _ H). rewrite unsigned_repr. auto. + generalize wordsize_pos wordsize_max_unsigned; omega. + assert (shl one y = repr (two_p uy)). + rewrite shl_mul_two_p. rewrite mul_commut. apply mul_one. + assert (and x (sub (shl one y) one) = modu x (repr (two_p uy))). + symmetry. rewrite H1. apply modu_and with (logn := y). + rewrite is_power2_two_p. unfold uy. rewrite repr_unsigned. auto. + omega. + rewrite H2. rewrite H1. + repeat rewrite shr_div_two_p. fold sx. fold uy. + assert (two_p uy > 0). apply two_p_gt_ZERO. omega. + assert (two_p uy < modulus). + rewrite modulus_power. apply two_p_monotone_strict. omega. + assert (two_p uy < half_modulus). + rewrite half_modulus_power. + apply two_p_monotone_strict. auto. + assert (two_p uy < modulus). + rewrite modulus_power. apply two_p_monotone_strict. omega. + assert (sub (repr (two_p uy)) one = repr (two_p uy - 1)). + unfold sub. apply eqm_samerepr. apply eqm_sub. apply eqm_sym; apply eqm_unsigned_repr. + rewrite unsigned_one. apply eqm_refl. + rewrite H7. rewrite add_signed. fold sx. + rewrite (signed_repr (two_p uy - 1)). rewrite signed_repr. + unfold modu. rewrite unsigned_repr. + unfold eq. rewrite unsigned_zero. rewrite unsigned_repr. + assert (unsigned x mod two_p uy = sx mod two_p uy). + apply eqmod_mod_eq; auto. apply eqmod_divides with modulus. + fold eqm. unfold sx. apply eqm_sym. apply eqm_signed_unsigned. + unfold modulus. rewrite two_power_nat_two_p. + exists (two_p (zwordsize - uy)). rewrite <- two_p_is_exp. + f_equal. fold zwordsize; omega. omega. omega. + rewrite H8. rewrite Zdiv_shift; auto. + unfold add. apply eqm_samerepr. apply eqm_add. + apply eqm_unsigned_repr. + destruct (zeq (sx mod two_p uy) 0); simpl. + rewrite unsigned_zero. apply eqm_refl. + rewrite unsigned_one. apply eqm_refl. + generalize (Z_mod_lt (unsigned x) (two_p uy) H3). unfold max_unsigned. omega. + unfold max_unsigned; omega. + generalize (signed_range x). fold sx. intros. split. omega. unfold max_signed. omega. + generalize min_signed_neg. unfold max_signed. omega. +Qed. + +(** Connections between [shr] and [shru]. *) + +Lemma shr_shru_positive: + forall x y, + signed x >= 0 -> + shr x y = shru x y. +Proof. + intros. + rewrite shr_div_two_p. rewrite shru_div_two_p. + rewrite signed_eq_unsigned. auto. apply signed_positive. auto. +Qed. + +Lemma and_positive: + forall x y, signed y >= 0 -> signed (and x y) >= 0. +Proof. + intros. + assert (unsigned y < half_modulus). rewrite signed_positive in H. unfold max_signed in H; omega. + generalize (sign_bit_of_unsigned y). rewrite zlt_true; auto. intros A. + generalize (sign_bit_of_unsigned (and x y)). rewrite bits_and. rewrite A. + rewrite andb_false_r. unfold signed. + destruct (zlt (unsigned (and x y)) half_modulus). + intros. generalize (unsigned_range (and x y)); omega. + congruence. + generalize wordsize_pos; omega. +Qed. + +Theorem shr_and_is_shru_and: + forall x y z, + lt y zero = false -> shr (and x y) z = shru (and x y) z. +Proof. + intros. apply shr_shru_positive. apply and_positive. + unfold lt in H. rewrite signed_zero in H. destruct (zlt (signed y) 0). congruence. auto. +Qed. + +(** ** Properties of integer zero extension and sign extension. *) + +Lemma Ziter_base: + forall (A: Type) n (f: A -> A) x, n <= 0 -> Z.iter n f x = x. +Proof. + intros. unfold Z.iter. destruct n; auto. compute in H. elim H; auto. +Qed. + +Lemma Ziter_succ: + forall (A: Type) n (f: A -> A) x, + 0 <= n -> Z.iter (Z.succ n) f x = f (Z.iter n f x). +Proof. + intros. destruct n; simpl. + - auto. + - rewrite Pos.add_1_r. apply Pos.iter_succ. + - compute in H. elim H; auto. +Qed. + +Lemma Znatlike_ind: + forall (P: Z -> Prop), + (forall n, n <= 0 -> P n) -> + (forall n, 0 <= n -> P n -> P (Z.succ n)) -> + forall n, P n. +Proof. + intros. destruct (zle 0 n). + apply natlike_ind; auto. apply H; omega. + apply H. omega. +Qed. + +Lemma Zzero_ext_spec: + forall n x i, 0 <= i -> + Z.testbit (Zzero_ext n x) i = if zlt i n then Z.testbit x i else false. +Proof. + unfold Zzero_ext. induction n using Znatlike_ind. + - intros. rewrite Ziter_base; auto. + rewrite zlt_false. rewrite Ztestbit_0; auto. omega. + - intros. rewrite Ziter_succ; auto. + rewrite Ztestbit_shiftin; auto. + rewrite (Ztestbit_eq i x); auto. + destruct (zeq i 0). + + subst i. rewrite zlt_true; auto. omega. + + rewrite IHn. destruct (zlt (Z.pred i) n). + rewrite zlt_true; auto. omega. + rewrite zlt_false; auto. omega. + omega. +Qed. + +Lemma bits_zero_ext: + forall n x i, 0 <= i -> + testbit (zero_ext n x) i = if zlt i n then testbit x i else false. +Proof. + intros. unfold zero_ext. destruct (zlt i zwordsize). + rewrite testbit_repr; auto. rewrite Zzero_ext_spec. auto. auto. + rewrite !bits_above; auto. destruct (zlt i n); auto. +Qed. + +Lemma Zsign_ext_spec: + forall n x i, 0 <= i -> 0 < n -> + Z.testbit (Zsign_ext n x) i = Z.testbit x (if zlt i n then i else n - 1). +Proof. + intros n0 x i I0 N0. + revert x i I0. pattern n0. apply Zlt_lower_bound_ind with (z := 1). + - unfold Zsign_ext. intros. + destruct (zeq x 1). + + subst x; simpl. + replace (if zlt i 1 then i else 0) with 0. + rewrite Ztestbit_base. + destruct (Z.odd x0). + apply Ztestbit_m1; auto. + apply Ztestbit_0. + destruct (zlt i 1); omega. + + set (x1 := Z.pred x). replace x1 with (Z.succ (Z.pred x1)). + rewrite Ziter_succ. rewrite Ztestbit_shiftin. + destruct (zeq i 0). + * subst i. rewrite zlt_true. rewrite Ztestbit_base; auto. omega. + * rewrite H. unfold x1. destruct (zlt (Z.pred i) (Z.pred x)). + rewrite zlt_true. rewrite (Ztestbit_eq i x0); auto. rewrite zeq_false; auto. omega. + rewrite zlt_false. rewrite (Ztestbit_eq (x - 1) x0). rewrite zeq_false; auto. + omega. omega. omega. unfold x1; omega. omega. + * omega. + * unfold x1; omega. + * omega. + - omega. +Qed. + +Lemma bits_sign_ext: + forall n x i, 0 <= i < zwordsize -> 0 < n -> + testbit (sign_ext n x) i = testbit x (if zlt i n then i else n - 1). +Proof. + intros. unfold sign_ext. + rewrite testbit_repr; auto. rewrite Zsign_ext_spec. destruct (zlt i n); auto. + omega. auto. +Qed. + +Hint Rewrite bits_zero_ext bits_sign_ext: ints. + +Theorem zero_ext_above: + forall n x, n >= zwordsize -> zero_ext n x = x. +Proof. + intros. apply same_bits_eq; intros. + rewrite bits_zero_ext. apply zlt_true. omega. omega. +Qed. + +Theorem sign_ext_above: + forall n x, n >= zwordsize -> sign_ext n x = x. +Proof. + intros. apply same_bits_eq; intros. + unfold sign_ext; rewrite testbit_repr; auto. + rewrite Zsign_ext_spec. rewrite zlt_true. auto. omega. omega. omega. +Qed. + +Theorem zero_ext_and: + forall n x, 0 <= n -> zero_ext n x = and x (repr (two_p n - 1)). +Proof. + bit_solve. rewrite testbit_repr; auto. rewrite Ztestbit_two_p_m1; intuition. + destruct (zlt i n). + rewrite andb_true_r; auto. + rewrite andb_false_r; auto. + tauto. +Qed. + +Theorem zero_ext_mod: + forall n x, 0 <= n < zwordsize -> + unsigned (zero_ext n x) = Zmod (unsigned x) (two_p n). +Proof. + intros. apply equal_same_bits. intros. + rewrite Ztestbit_mod_two_p; auto. + fold (testbit (zero_ext n x) i). + destruct (zlt i zwordsize). + rewrite bits_zero_ext; auto. + rewrite bits_above. rewrite zlt_false; auto. omega. omega. + omega. +Qed. + +Theorem zero_ext_widen: + forall x n n', 0 <= n <= n' -> + zero_ext n' (zero_ext n x) = zero_ext n x. +Proof. + bit_solve. destruct (zlt i n). + apply zlt_true. omega. + destruct (zlt i n'); auto. + tauto. tauto. +Qed. + +Theorem sign_ext_widen: + forall x n n', 0 < n <= n' -> + sign_ext n' (sign_ext n x) = sign_ext n x. +Proof. + intros. destruct (zlt n' zwordsize). + bit_solve. destruct (zlt i n'). + auto. + rewrite (zlt_false _ i n). + destruct (zlt (n' - 1) n); f_equal; omega. + omega. omega. + destruct (zlt i n'); omega. + omega. omega. + apply sign_ext_above; auto. +Qed. + +Theorem sign_zero_ext_widen: + forall x n n', 0 <= n < n' -> + sign_ext n' (zero_ext n x) = zero_ext n x. +Proof. + intros. destruct (zlt n' zwordsize). + bit_solve. + destruct (zlt i n'). + auto. + rewrite !zlt_false. auto. omega. omega. omega. + destruct (zlt i n'); omega. + omega. + apply sign_ext_above; auto. +Qed. + +Theorem zero_ext_narrow: + forall x n n', 0 <= n <= n' -> + zero_ext n (zero_ext n' x) = zero_ext n x. +Proof. + bit_solve. destruct (zlt i n). + apply zlt_true. omega. + auto. + omega. omega. omega. +Qed. + +Theorem sign_ext_narrow: + forall x n n', 0 < n <= n' -> + sign_ext n (sign_ext n' x) = sign_ext n x. +Proof. + intros. destruct (zlt n zwordsize). + bit_solve. destruct (zlt i n); f_equal; apply zlt_true; omega. + omega. + destruct (zlt i n); omega. + omega. omega. + rewrite (sign_ext_above n'). auto. omega. +Qed. + +Theorem zero_sign_ext_narrow: + forall x n n', 0 < n <= n' -> + zero_ext n (sign_ext n' x) = zero_ext n x. +Proof. + intros. destruct (zlt n' zwordsize). + bit_solve. + destruct (zlt i n); auto. + rewrite zlt_true; auto. omega. + omega. omega. omega. + rewrite sign_ext_above; auto. +Qed. + +Theorem zero_ext_idem: + forall n x, 0 <= n -> zero_ext n (zero_ext n x) = zero_ext n x. +Proof. + intros. apply zero_ext_widen. omega. +Qed. + +Theorem sign_ext_idem: + forall n x, 0 < n -> sign_ext n (sign_ext n x) = sign_ext n x. +Proof. + intros. apply sign_ext_widen. omega. +Qed. + +Theorem sign_ext_zero_ext: + forall n x, 0 < n -> sign_ext n (zero_ext n x) = sign_ext n x. +Proof. + intros. destruct (zlt n zwordsize). + bit_solve. + destruct (zlt i n). + rewrite zlt_true; auto. + rewrite zlt_true; auto. omega. + destruct (zlt i n); omega. + rewrite zero_ext_above; auto. +Qed. + +Theorem zero_ext_sign_ext: + forall n x, 0 < n -> zero_ext n (sign_ext n x) = zero_ext n x. +Proof. + intros. apply zero_sign_ext_narrow. omega. +Qed. + +Theorem sign_ext_equal_if_zero_equal: + forall n x y, 0 < n -> + zero_ext n x = zero_ext n y -> + sign_ext n x = sign_ext n y. +Proof. + intros. rewrite <- (sign_ext_zero_ext n x H). + rewrite <- (sign_ext_zero_ext n y H). congruence. +Qed. + +Theorem zero_ext_shru_shl: + forall n x, + 0 < n < zwordsize -> + let y := repr (zwordsize - n) in + zero_ext n x = shru (shl x y) y. +Proof. + intros. + assert (unsigned y = zwordsize - n). + unfold y. apply unsigned_repr. generalize wordsize_max_unsigned. omega. + apply same_bits_eq; intros. + rewrite bits_zero_ext. + rewrite bits_shru; auto. + destruct (zlt i n). + rewrite zlt_true. rewrite bits_shl. rewrite zlt_false. f_equal. omega. + omega. omega. omega. + rewrite zlt_false. auto. omega. + omega. +Qed. + +Theorem sign_ext_shr_shl: + forall n x, + 0 < n < zwordsize -> + let y := repr (zwordsize - n) in + sign_ext n x = shr (shl x y) y. +Proof. + intros. + assert (unsigned y = zwordsize - n). + unfold y. apply unsigned_repr. generalize wordsize_max_unsigned. omega. + apply same_bits_eq; intros. + rewrite bits_sign_ext. + rewrite bits_shr; auto. + destruct (zlt i n). + rewrite zlt_true. rewrite bits_shl. rewrite zlt_false. f_equal. omega. + omega. omega. omega. + rewrite zlt_false. rewrite bits_shl. rewrite zlt_false. f_equal. omega. + omega. omega. omega. omega. omega. +Qed. + +(** [zero_ext n x] is the unique integer congruent to [x] modulo [2^n] + in the range [0...2^n-1]. *) + +Lemma zero_ext_range: + forall n x, 0 <= n < zwordsize -> 0 <= unsigned (zero_ext n x) < two_p n. +Proof. + intros. rewrite zero_ext_mod; auto. apply Z_mod_lt. apply two_p_gt_ZERO. omega. +Qed. + +Lemma eqmod_zero_ext: + forall n x, 0 <= n < zwordsize -> eqmod (two_p n) (unsigned (zero_ext n x)) (unsigned x). +Proof. + intros. rewrite zero_ext_mod; auto. apply eqmod_sym. apply eqmod_mod. + apply two_p_gt_ZERO. omega. +Qed. + +(** [sign_ext n x] is the unique integer congruent to [x] modulo [2^n] + in the range [-2^(n-1)...2^(n-1) - 1]. *) + +Lemma sign_ext_range: + forall n x, 0 < n < zwordsize -> -two_p (n-1) <= signed (sign_ext n x) < two_p (n-1). +Proof. + intros. rewrite sign_ext_shr_shl; auto. + set (X := shl x (repr (zwordsize - n))). + assert (two_p (n - 1) > 0) by (apply two_p_gt_ZERO; omega). + assert (unsigned (repr (zwordsize - n)) = zwordsize - n). + apply unsigned_repr. + split. omega. generalize wordsize_max_unsigned; omega. + rewrite shr_div_two_p. + rewrite signed_repr. + rewrite H1. + apply Zdiv_interval_1. + omega. omega. apply two_p_gt_ZERO; omega. + replace (- two_p (n - 1) * two_p (zwordsize - n)) + with (- (two_p (n - 1) * two_p (zwordsize - n))) by ring. + rewrite <- two_p_is_exp. + replace (n - 1 + (zwordsize - n)) with (zwordsize - 1) by omega. + rewrite <- half_modulus_power. + generalize (signed_range X). unfold min_signed, max_signed. omega. + omega. omega. + apply Zdiv_interval_2. apply signed_range. + generalize min_signed_neg; omega. + generalize max_signed_pos; omega. + rewrite H1. apply two_p_gt_ZERO. omega. +Qed. + +Lemma eqmod_sign_ext': + forall n x, 0 < n < zwordsize -> + eqmod (two_p n) (unsigned (sign_ext n x)) (unsigned x). +Proof. + intros. + set (N := Z.to_nat n). + assert (Z.of_nat N = n) by (apply Z2Nat.id; omega). + rewrite <- H0. rewrite <- two_power_nat_two_p. + apply eqmod_same_bits; intros. + rewrite H0 in H1. rewrite H0. + fold (testbit (sign_ext n x) i). rewrite bits_sign_ext. + rewrite zlt_true. auto. omega. omega. omega. +Qed. + +Lemma eqmod_sign_ext: + forall n x, 0 < n < zwordsize -> + eqmod (two_p n) (signed (sign_ext n x)) (unsigned x). +Proof. + intros. apply eqmod_trans with (unsigned (sign_ext n x)). + apply eqmod_divides with modulus. apply eqm_signed_unsigned. + exists (two_p (zwordsize - n)). + unfold modulus. rewrite two_power_nat_two_p. fold zwordsize. + rewrite <- two_p_is_exp. f_equal. omega. omega. omega. + apply eqmod_sign_ext'; auto. +Qed. + +(** ** Properties of [one_bits] (decomposition in sum of powers of two) *) + +Theorem one_bits_range: + forall x i, In i (one_bits x) -> ltu i iwordsize = true. +Proof. + assert (A: forall p, 0 <= p < zwordsize -> ltu (repr p) iwordsize = true). + intros. unfold ltu, iwordsize. apply zlt_true. + repeat rewrite unsigned_repr. tauto. + generalize wordsize_max_unsigned; omega. + generalize wordsize_max_unsigned; omega. + unfold one_bits. intros. + destruct (list_in_map_inv _ _ _ H) as [i0 [EQ IN]]. + subst i. apply A. apply Z_one_bits_range with (unsigned x); auto. +Qed. + +Fixpoint int_of_one_bits (l: list int) : int := + match l with + | nil => zero + | a :: b => add (shl one a) (int_of_one_bits b) + end. + +Theorem one_bits_decomp: + forall x, x = int_of_one_bits (one_bits x). +Proof. + intros. + transitivity (repr (powerserie (Z_one_bits wordsize (unsigned x) 0))). + transitivity (repr (unsigned x)). + auto with ints. decEq. apply Z_one_bits_powerserie. + auto with ints. + unfold one_bits. + generalize (Z_one_bits_range (unsigned x)). + generalize (Z_one_bits wordsize (unsigned x) 0). + induction l. + intros; reflexivity. + intros; simpl. rewrite <- IHl. unfold add. apply eqm_samerepr. + apply eqm_add. rewrite shl_mul_two_p. rewrite mul_commut. + rewrite mul_one. apply eqm_unsigned_repr_r. + rewrite unsigned_repr. auto with ints. + generalize (H a (in_eq _ _)). generalize wordsize_max_unsigned. omega. + auto with ints. + intros; apply H; auto with coqlib. +Qed. + +(** ** Properties of comparisons *) + +Theorem negate_cmp: + forall c x y, cmp (negate_comparison c) x y = negb (cmp c x y). +Proof. + intros. destruct c; simpl; try rewrite negb_elim; auto. +Qed. + +Theorem negate_cmpu: + forall c x y, cmpu (negate_comparison c) x y = negb (cmpu c x y). +Proof. + intros. destruct c; simpl; try rewrite negb_elim; auto. +Qed. + +Theorem swap_cmp: + forall c x y, cmp (swap_comparison c) x y = cmp c y x. +Proof. + intros. destruct c; simpl; auto. apply eq_sym. decEq. apply eq_sym. +Qed. + +Theorem swap_cmpu: + forall c x y, cmpu (swap_comparison c) x y = cmpu c y x. +Proof. + intros. destruct c; simpl; auto. apply eq_sym. decEq. apply eq_sym. +Qed. + +Lemma translate_eq: + forall x y d, + eq (add x d) (add y d) = eq x y. +Proof. + intros. unfold eq. case (zeq (unsigned x) (unsigned y)); intro. + unfold add. rewrite e. apply zeq_true. + apply zeq_false. unfold add. red; intro. apply n. + apply eqm_small_eq; auto with ints. + replace (unsigned x) with ((unsigned x + unsigned d) - unsigned d). + replace (unsigned y) with ((unsigned y + unsigned d) - unsigned d). + apply eqm_sub. apply eqm_trans with (unsigned (repr (unsigned x + unsigned d))). + eauto with ints. apply eqm_trans with (unsigned (repr (unsigned y + unsigned d))). + eauto with ints. eauto with ints. eauto with ints. + omega. omega. +Qed. + +Lemma translate_ltu: + forall x y d, + 0 <= unsigned x + unsigned d <= max_unsigned -> + 0 <= unsigned y + unsigned d <= max_unsigned -> + ltu (add x d) (add y d) = ltu x y. +Proof. + intros. unfold add. unfold ltu. + repeat rewrite unsigned_repr; auto. case (zlt (unsigned x) (unsigned y)); intro. + apply zlt_true. omega. + apply zlt_false. omega. +Qed. + +Theorem translate_cmpu: + forall c x y d, + 0 <= unsigned x + unsigned d <= max_unsigned -> + 0 <= unsigned y + unsigned d <= max_unsigned -> + cmpu c (add x d) (add y d) = cmpu c x y. +Proof. + intros. unfold cmpu. + rewrite translate_eq. repeat rewrite translate_ltu; auto. +Qed. + +Lemma translate_lt: + forall x y d, + min_signed <= signed x + signed d <= max_signed -> + min_signed <= signed y + signed d <= max_signed -> + lt (add x d) (add y d) = lt x y. +Proof. + intros. repeat rewrite add_signed. unfold lt. + repeat rewrite signed_repr; auto. case (zlt (signed x) (signed y)); intro. + apply zlt_true. omega. + apply zlt_false. omega. +Qed. + +Theorem translate_cmp: + forall c x y d, + min_signed <= signed x + signed d <= max_signed -> + min_signed <= signed y + signed d <= max_signed -> + cmp c (add x d) (add y d) = cmp c x y. +Proof. + intros. unfold cmp. + rewrite translate_eq. repeat rewrite translate_lt; auto. +Qed. + +Theorem notbool_isfalse_istrue: + forall x, is_false x -> is_true (notbool x). +Proof. + unfold is_false, is_true, notbool; intros; subst x. + rewrite eq_true. apply one_not_zero. +Qed. + +Theorem notbool_istrue_isfalse: + forall x, is_true x -> is_false (notbool x). +Proof. + unfold is_false, is_true, notbool; intros. + generalize (eq_spec x zero). case (eq x zero); intro. + contradiction. auto. +Qed. + +Theorem ltu_range_test: + forall x y, + ltu x y = true -> unsigned y <= max_signed -> + 0 <= signed x < unsigned y. +Proof. + intros. + unfold ltu in H. destruct (zlt (unsigned x) (unsigned y)); try discriminate. + rewrite signed_eq_unsigned. + generalize (unsigned_range x). omega. omega. +Qed. + +Theorem lt_sub_overflow: + forall x y, + xor (sub_overflow x y zero) (negative (sub x y)) = if lt x y then one else zero. +Proof. + intros. unfold negative, sub_overflow, lt. rewrite sub_signed. + rewrite signed_zero. rewrite Zminus_0_r. + generalize (signed_range x) (signed_range y). + set (X := signed x); set (Y := signed y). intros RX RY. + unfold min_signed, max_signed in *. + generalize half_modulus_pos half_modulus_modulus; intros HM MM. + destruct (zle 0 (X - Y)). +- unfold proj_sumbool at 1; rewrite zle_true at 1 by omega. simpl. + rewrite (zlt_false _ X) by omega. + destruct (zlt (X - Y) half_modulus). + + unfold proj_sumbool; rewrite zle_true by omega. + rewrite signed_repr. rewrite zlt_false by omega. apply xor_idem. + unfold min_signed, max_signed; omega. + + unfold proj_sumbool; rewrite zle_false by omega. + replace (signed (repr (X - Y))) with (X - Y - modulus). + rewrite zlt_true by omega. apply xor_idem. + rewrite signed_repr_eq. replace ((X - Y) mod modulus) with (X - Y). + rewrite zlt_false; auto. + symmetry. apply Zmod_unique with 0; omega. +- unfold proj_sumbool at 2. rewrite zle_true at 1 by omega. rewrite andb_true_r. + rewrite (zlt_true _ X) by omega. + destruct (zlt (X - Y) (-half_modulus)). + + unfold proj_sumbool; rewrite zle_false by omega. + replace (signed (repr (X - Y))) with (X - Y + modulus). + rewrite zlt_false by omega. apply xor_zero. + rewrite signed_repr_eq. replace ((X - Y) mod modulus) with (X - Y + modulus). + rewrite zlt_true by omega; auto. + symmetry. apply Zmod_unique with (-1); omega. + + unfold proj_sumbool; rewrite zle_true by omega. + rewrite signed_repr. rewrite zlt_true by omega. apply xor_zero_l. + unfold min_signed, max_signed; omega. +Qed. + +Lemma signed_eq: + forall x y, eq x y = zeq (signed x) (signed y). +Proof. + intros. unfold eq. unfold proj_sumbool. + destruct (zeq (unsigned x) (unsigned y)); + destruct (zeq (signed x) (signed y)); auto. + elim n. unfold signed. rewrite e; auto. + elim n. apply eqm_small_eq; auto with ints. + eapply eqm_trans. apply eqm_sym. apply eqm_signed_unsigned. + rewrite e. apply eqm_signed_unsigned. +Qed. + +Lemma not_lt: + forall x y, negb (lt y x) = (lt x y || eq x y). +Proof. + intros. unfold lt. rewrite signed_eq. unfold proj_sumbool. + destruct (zlt (signed y) (signed x)). + rewrite zlt_false. rewrite zeq_false. auto. omega. omega. + destruct (zeq (signed x) (signed y)). + rewrite zlt_false. auto. omega. + rewrite zlt_true. auto. omega. +Qed. + +Lemma lt_not: + forall x y, lt y x = negb (lt x y) && negb (eq x y). +Proof. + intros. rewrite <- negb_orb. rewrite <- not_lt. rewrite negb_involutive. auto. +Qed. + +Lemma not_ltu: + forall x y, negb (ltu y x) = (ltu x y || eq x y). +Proof. + intros. unfold ltu, eq. + destruct (zlt (unsigned y) (unsigned x)). + rewrite zlt_false. rewrite zeq_false. auto. omega. omega. + destruct (zeq (unsigned x) (unsigned y)). + rewrite zlt_false. auto. omega. + rewrite zlt_true. auto. omega. +Qed. + +Lemma ltu_not: + forall x y, ltu y x = negb (ltu x y) && negb (eq x y). +Proof. + intros. rewrite <- negb_orb. rewrite <- not_ltu. rewrite negb_involutive. auto. +Qed. + + +(** Non-overlapping test *) + +Definition no_overlap (ofs1: int) (sz1: Z) (ofs2: int) (sz2: Z) : bool := + let x1 := unsigned ofs1 in let x2 := unsigned ofs2 in + zlt (x1 + sz1) modulus && zlt (x2 + sz2) modulus + && (zle (x1 + sz1) x2 || zle (x2 + sz2) x1). + +Lemma no_overlap_sound: + forall ofs1 sz1 ofs2 sz2 base, + sz1 > 0 -> sz2 > 0 -> no_overlap ofs1 sz1 ofs2 sz2 = true -> + unsigned (add base ofs1) + sz1 <= unsigned (add base ofs2) + \/ unsigned (add base ofs2) + sz2 <= unsigned (add base ofs1). +Proof. + intros. + destruct (andb_prop _ _ H1). clear H1. + destruct (andb_prop _ _ H2). clear H2. + apply proj_sumbool_true in H1. + apply proj_sumbool_true in H4. + assert (unsigned ofs1 + sz1 <= unsigned ofs2 \/ unsigned ofs2 + sz2 <= unsigned ofs1). + destruct (orb_prop _ _ H3). left. eapply proj_sumbool_true; eauto. right. eapply proj_sumbool_true; eauto. + clear H3. + generalize (unsigned_range ofs1) (unsigned_range ofs2). intros P Q. + generalize (unsigned_add_either base ofs1) (unsigned_add_either base ofs2). + intros [C|C] [D|D]; omega. +Qed. + +(** Size of integers, in bits. *) + +Definition Zsize (x: Z) : Z := + match x with + | Zpos p => Zpos (Pos.size p) + | _ => 0 + end. + +Definition size (x: int) : Z := Zsize (unsigned x). + +Remark Zsize_pos: forall x, 0 <= Zsize x. +Proof. + destruct x; simpl. omega. compute; intuition congruence. omega. +Qed. + +Remark Zsize_pos': forall x, 0 < x -> 0 < Zsize x. +Proof. + destruct x; simpl; intros; try discriminate. compute; auto. +Qed. + +Lemma Zsize_shiftin: + forall b x, 0 < x -> Zsize (Zshiftin b x) = Zsucc (Zsize x). +Proof. + intros. destruct x; compute in H; try discriminate. + destruct b. + change (Zshiftin true (Zpos p)) with (Zpos (p~1)). + simpl. f_equal. rewrite Pos.add_1_r; auto. + change (Zshiftin false (Zpos p)) with (Zpos (p~0)). + simpl. f_equal. rewrite Pos.add_1_r; auto. +Qed. + +Lemma Ztestbit_size_1: + forall x, 0 < x -> Z.testbit x (Zpred (Zsize x)) = true. +Proof. + intros x0 POS0; pattern x0; apply Zshiftin_pos_ind; auto. + intros. rewrite Zsize_shiftin; auto. + replace (Z.pred (Z.succ (Zsize x))) with (Z.succ (Z.pred (Zsize x))) by omega. + rewrite Ztestbit_shiftin_succ. auto. generalize (Zsize_pos' x H); omega. +Qed. + +Lemma Ztestbit_size_2: + forall x, 0 <= x -> forall i, i >= Zsize x -> Z.testbit x i = false. +Proof. + intros x0 POS0. destruct (zeq x0 0). + - subst x0; intros. apply Ztestbit_0. + - pattern x0; apply Zshiftin_pos_ind. + + simpl. intros. change 1 with (Zshiftin true 0). rewrite Ztestbit_shiftin. + rewrite zeq_false. apply Ztestbit_0. omega. omega. + + intros. rewrite Zsize_shiftin in H1; auto. + generalize (Zsize_pos' _ H); intros. + rewrite Ztestbit_shiftin. rewrite zeq_false. apply H0. omega. + omega. omega. + + omega. +Qed. + +Lemma Zsize_interval_1: + forall x, 0 <= x -> 0 <= x < two_p (Zsize x). +Proof. + intros. + assert (x = x mod (two_p (Zsize x))). + apply equal_same_bits; intros. + rewrite Ztestbit_mod_two_p; auto. + destruct (zlt i (Zsize x)). auto. apply Ztestbit_size_2; auto. + apply Zsize_pos; auto. + rewrite H0 at 1. rewrite H0 at 3. apply Z_mod_lt. apply two_p_gt_ZERO. apply Zsize_pos; auto. +Qed. + +Lemma Zsize_interval_2: + forall x n, 0 <= n -> 0 <= x < two_p n -> n >= Zsize x. +Proof. + intros. set (N := Z.to_nat n). + assert (Z.of_nat N = n) by (apply Z2Nat.id; auto). + rewrite <- H1 in H0. rewrite <- two_power_nat_two_p in H0. + destruct (zeq x 0). + subst x; simpl; omega. + destruct (zlt n (Zsize x)); auto. + exploit (Ztestbit_above N x (Zpred (Zsize x))). auto. omega. + rewrite Ztestbit_size_1. congruence. omega. +Qed. + +Lemma Zsize_monotone: + forall x y, 0 <= x <= y -> Zsize x <= Zsize y. +Proof. + intros. apply Zge_le. apply Zsize_interval_2. apply Zsize_pos. + exploit (Zsize_interval_1 y). omega. + omega. +Qed. + +Theorem size_zero: size zero = 0. +Proof. + unfold size; rewrite unsigned_zero; auto. +Qed. + +Theorem bits_size_1: + forall x, x = zero \/ testbit x (Zpred (size x)) = true. +Proof. + intros. destruct (zeq (unsigned x) 0). + left. rewrite <- (repr_unsigned x). rewrite e; auto. + right. apply Ztestbit_size_1. generalize (unsigned_range x); omega. +Qed. + +Theorem bits_size_2: + forall x i, size x <= i -> testbit x i = false. +Proof. + intros. apply Ztestbit_size_2. generalize (unsigned_range x); omega. + fold (size x); omega. +Qed. + +Theorem size_range: + forall x, 0 <= size x <= zwordsize. +Proof. + intros; split. apply Zsize_pos. + destruct (bits_size_1 x). + subst x; unfold size; rewrite unsigned_zero; simpl. generalize wordsize_pos; omega. + destruct (zle (size x) zwordsize); auto. + rewrite bits_above in H. congruence. omega. +Qed. + +Theorem bits_size_3: + forall x n, + 0 <= n -> + (forall i, n <= i < zwordsize -> testbit x i = false) -> + size x <= n. +Proof. + intros. destruct (zle (size x) n). auto. + destruct (bits_size_1 x). + subst x. unfold size; rewrite unsigned_zero; assumption. + rewrite (H0 (Z.pred (size x))) in H1. congruence. + generalize (size_range x); omega. +Qed. + +Theorem bits_size_4: + forall x n, + 0 <= n -> + testbit x (Zpred n) = true -> + (forall i, n <= i < zwordsize -> testbit x i = false) -> + size x = n. +Proof. + intros. + assert (size x <= n). + apply bits_size_3; auto. + destruct (zlt (size x) n). + rewrite bits_size_2 in H0. congruence. omega. + omega. +Qed. + +Theorem size_interval_1: + forall x, 0 <= unsigned x < two_p (size x). +Proof. + intros; apply Zsize_interval_1. generalize (unsigned_range x); omega. +Qed. + +Theorem size_interval_2: + forall x n, 0 <= n -> 0 <= unsigned x < two_p n -> n >= size x. +Proof. + intros. apply Zsize_interval_2; auto. +Qed. + +Theorem size_and: + forall a b, size (and a b) <= Z.min (size a) (size b). +Proof. + intros. + assert (0 <= Z.min (size a) (size b)). + generalize (size_range a) (size_range b). zify; omega. + apply bits_size_3. auto. intros. + rewrite bits_and. zify. subst z z0. destruct H1. + rewrite (bits_size_2 a). auto. omega. + rewrite (bits_size_2 b). apply andb_false_r. omega. + omega. +Qed. + +Corollary and_interval: + forall a b, 0 <= unsigned (and a b) < two_p (Z.min (size a) (size b)). +Proof. + intros. + generalize (size_interval_1 (and a b)); intros. + assert (two_p (size (and a b)) <= two_p (Z.min (size a) (size b))). + apply two_p_monotone. split. generalize (size_range (and a b)); omega. + apply size_and. + omega. +Qed. + +Theorem size_or: + forall a b, size (or a b) = Z.max (size a) (size b). +Proof. + intros. generalize (size_range a) (size_range b); intros. + destruct (bits_size_1 a). + subst a. rewrite size_zero. rewrite or_zero_l. zify; omega. + destruct (bits_size_1 b). + subst b. rewrite size_zero. rewrite or_zero. zify; omega. + zify. destruct H3 as [[P Q] | [P Q]]; subst. + apply bits_size_4. tauto. rewrite bits_or. rewrite H2. apply orb_true_r. + omega. + intros. rewrite bits_or. rewrite !bits_size_2. auto. omega. omega. omega. + apply bits_size_4. tauto. rewrite bits_or. rewrite H1. apply orb_true_l. + destruct (zeq (size a) 0). unfold testbit in H1. rewrite Z.testbit_neg_r in H1. + congruence. omega. omega. + intros. rewrite bits_or. rewrite !bits_size_2. auto. omega. omega. omega. +Qed. + +Corollary or_interval: + forall a b, 0 <= unsigned (or a b) < two_p (Z.max (size a) (size b)). +Proof. + intros. rewrite <- size_or. apply size_interval_1. +Qed. + +Theorem size_xor: + forall a b, size (xor a b) <= Z.max (size a) (size b). +Proof. + intros. + assert (0 <= Z.max (size a) (size b)). + generalize (size_range a) (size_range b). zify; omega. + apply bits_size_3. auto. intros. + rewrite bits_xor. rewrite !bits_size_2. auto. + zify; omega. + zify; omega. + omega. +Qed. + +Corollary xor_interval: + forall a b, 0 <= unsigned (xor a b) < two_p (Z.max (size a) (size b)). +Proof. + intros. + generalize (size_interval_1 (xor a b)); intros. + assert (two_p (size (xor a b)) <= two_p (Z.max (size a) (size b))). + apply two_p_monotone. split. generalize (size_range (xor a b)); omega. + apply size_xor. + omega. +Qed. + +End Make. + +(** * Specialization to integers of size 8, 32, and 64 bits *) + +Module Wordsize_32. + Definition wordsize := 32%nat. + Remark wordsize_not_zero: wordsize <> 0%nat. + Proof. unfold wordsize; congruence. Qed. +End Wordsize_32. + +Strategy opaque [Wordsize_32.wordsize]. + +Module Int := Make(Wordsize_32). + +Strategy 0 [Wordsize_32.wordsize]. + +Notation int := Int.int. + +Remark int_wordsize_divides_modulus: + Zdivide (Z_of_nat Int.wordsize) Int.modulus. +Proof. + exists (two_p (32-5)); reflexivity. +Qed. + +Module Wordsize_256. + Definition wordsize := 256%nat. + Remark wordsize_not_zero: wordsize <> 0%nat. + Proof. unfold wordsize; congruence. Qed. +End Wordsize_256. + +Strategy opaque [Wordsize_256.wordsize]. + +Module Int256 := Make(Wordsize_256). + +Strategy 0 [Wordsize_256.wordsize]. + +Notation int256 := Int256.int. + +Module Wordsize_8. + Definition wordsize := 8%nat. + Remark wordsize_not_zero: wordsize <> 0%nat. + Proof. unfold wordsize; congruence. Qed. +End Wordsize_8. + +Strategy opaque [Wordsize_8.wordsize]. + +Module Byte := Make(Wordsize_8). + +Strategy 0 [Wordsize_8.wordsize]. + +Notation byte := Byte.int. + +Module Wordsize_64. + Definition wordsize := 64%nat. + Remark wordsize_not_zero: wordsize <> 0%nat. + Proof. unfold wordsize; congruence. Qed. +End Wordsize_64. + +Strategy opaque [Wordsize_64.wordsize]. + +Module Int64. + +Include Make(Wordsize_64). + +(** Shifts with amount given as a 32-bit integer *) + +Definition iwordsize': Int.int := Int.repr zwordsize. + +Definition shl' (x: int) (y: Int.int): int := + repr (Z.shiftl (unsigned x) (Int.unsigned y)). +Definition shru' (x: int) (y: Int.int): int := + repr (Z.shiftr (unsigned x) (Int.unsigned y)). +Definition shr' (x: int) (y: Int.int): int := + repr (Z.shiftr (signed x) (Int.unsigned y)). +Definition rol' (x: int) (y: Int.int): int := + rol x (repr (Int.unsigned y)). +Definition shrx' (x: int) (y: Int.int): int := + divs x (shl' one y). +Definition shr_carry' (x: int) (y: Int.int): int := + if lt x zero && negb (eq (and x (sub (shl' one y) one)) zero) + then one else zero. + +Lemma bits_shl': + forall x y i, + 0 <= i < zwordsize -> + testbit (shl' x y) i = + if zlt i (Int.unsigned y) then false else testbit x (i - Int.unsigned y). +Proof. + intros. unfold shl'. rewrite testbit_repr; auto. + destruct (zlt i (Int.unsigned y)). + apply Z.shiftl_spec_low. auto. + apply Z.shiftl_spec_high. omega. omega. +Qed. + +Lemma bits_shru': + forall x y i, + 0 <= i < zwordsize -> + testbit (shru' x y) i = + if zlt (i + Int.unsigned y) zwordsize then testbit x (i + Int.unsigned y) else false. +Proof. + intros. unfold shru'. rewrite testbit_repr; auto. + rewrite Z.shiftr_spec. fold (testbit x (i + Int.unsigned y)). + destruct (zlt (i + Int.unsigned y) zwordsize). + auto. + apply bits_above; auto. + omega. +Qed. + +Lemma bits_shr': + forall x y i, + 0 <= i < zwordsize -> + testbit (shr' x y) i = + testbit x (if zlt (i + Int.unsigned y) zwordsize then i + Int.unsigned y else zwordsize - 1). +Proof. + intros. unfold shr'. rewrite testbit_repr; auto. + rewrite Z.shiftr_spec. apply bits_signed. + generalize (Int.unsigned_range y); omega. + omega. +Qed. + +Lemma shl'_mul_two_p: + forall x y, + shl' x y = mul x (repr (two_p (Int.unsigned y))). +Proof. + intros. unfold shl', mul. apply eqm_samerepr. + rewrite Zshiftl_mul_two_p. apply eqm_mult. apply eqm_refl. apply eqm_unsigned_repr. + generalize (Int.unsigned_range y); omega. +Qed. + +Lemma shl'_one_two_p: + forall y, shl' one y = repr (two_p (Int.unsigned y)). +Proof. + intros. rewrite shl'_mul_two_p. rewrite mul_commut. rewrite mul_one. auto. +Qed. + +Theorem shl'_mul: + forall x y, + shl' x y = mul x (shl' one y). +Proof. + intros. rewrite shl'_one_two_p. apply shl'_mul_two_p. +Qed. + +Theorem shl'_zero: + forall x, shl' x Int.zero = x. +Proof. + intros. unfold shl'. rewrite Int.unsigned_zero. unfold Z.shiftl. + apply repr_unsigned. +Qed. + +Theorem shru'_zero : + forall x, shru' x Int.zero = x. +Proof. + intros. unfold shru'. rewrite Int.unsigned_zero. unfold Z.shiftr. + apply repr_unsigned. +Qed. + +Theorem shr'_zero : + forall x, shr' x Int.zero = x. +Proof. + intros. unfold shr'. rewrite Int.unsigned_zero. unfold Z.shiftr. + apply repr_signed. +Qed. + +Theorem shrx'_zero: + forall x, shrx' x Int.zero = x. +Proof. + intros. change (shrx' x Int.zero) with (shrx x zero). apply shrx_zero. compute; auto. +Qed. + +Theorem shrx'_carry: + forall x y, + Int.ltu y (Int.repr 63) = true -> + shrx' x y = add (shr' x y) (shr_carry' x y). +Proof. + intros. apply Int.ltu_inv in H. change (Int.unsigned (Int.repr 63)) with 63 in H. + set (y1 := Int64.repr (Int.unsigned y)). + assert (U: unsigned y1 = Int.unsigned y). + { apply unsigned_repr. assert (63 < max_unsigned) by reflexivity. omega. } + transitivity (shrx x y1). +- unfold shrx', shrx, shl', shl. rewrite U; auto. +- rewrite shrx_carry. ++ f_equal. + unfold shr, shr'. rewrite U; auto. + unfold shr_carry, shr_carry', shl, shl'. rewrite U; auto. ++ unfold ltu. apply zlt_true. rewrite U; tauto. +Qed. + +Theorem shrx'_shr_2: + forall x y, + Int.ltu y (Int.repr 63) = true -> + shrx' x y = shr' (add x (shru' (shr' x (Int.repr 63)) (Int.sub (Int.repr 64) y))) y. +Proof. + intros. + set (z := repr (Int.unsigned y)). + apply Int.ltu_inv in H. change (Int.unsigned (Int.repr 63)) with 63 in H. + assert (N1: 63 < max_unsigned) by reflexivity. + assert (N2: 63 < Int.max_unsigned) by reflexivity. + assert (A: unsigned z = Int.unsigned y). + { unfold z; apply unsigned_repr; omega. } + assert (B: unsigned (sub (repr 64) z) = Int.unsigned (Int.sub (Int.repr 64) y)). + { unfold z. unfold sub, Int.sub. + change (unsigned (repr 64)) with 64. + change (Int.unsigned (Int.repr 64)) with 64. + rewrite (unsigned_repr (Int.unsigned y)) by omega. + rewrite unsigned_repr, Int.unsigned_repr by omega. + auto. } + unfold shrx', shr', shru', shl'. + rewrite <- A. + change (Int.unsigned (Int.repr 63)) with (unsigned (repr 63)). + rewrite <- B. + apply shrx_shr_2. + unfold ltu. apply zlt_true. change (unsigned z < 63). rewrite A; omega. +Qed. + +Remark int_ltu_2_inv: + forall y z, + Int.ltu y iwordsize' = true -> + Int.ltu z iwordsize' = true -> + Int.unsigned (Int.add y z) <= Int.unsigned iwordsize' -> + let y' := repr (Int.unsigned y) in + let z' := repr (Int.unsigned z) in + Int.unsigned y = unsigned y' + /\ Int.unsigned z = unsigned z' + /\ ltu y' iwordsize = true + /\ ltu z' iwordsize = true + /\ Int.unsigned (Int.add y z) = unsigned (add y' z') + /\ add y' z' = repr (Int.unsigned (Int.add y z)). +Proof. + intros. apply Int.ltu_inv in H. apply Int.ltu_inv in H0. + change (Int.unsigned iwordsize') with 64 in *. + assert (128 < max_unsigned) by reflexivity. + assert (128 < Int.max_unsigned) by reflexivity. + assert (Y: unsigned y' = Int.unsigned y) by (apply unsigned_repr; omega). + assert (Z: unsigned z' = Int.unsigned z) by (apply unsigned_repr; omega). + assert (P: Int.unsigned (Int.add y z) = unsigned (add y' z')). + { unfold Int.add. rewrite Int.unsigned_repr by omega. + unfold add. rewrite unsigned_repr by omega. congruence. } + intuition auto. + apply zlt_true. rewrite Y; auto. + apply zlt_true. rewrite Z; auto. + rewrite P. rewrite repr_unsigned. auto. +Qed. + +Theorem or_ror': + forall x y z, + Int.ltu y iwordsize' = true -> + Int.ltu z iwordsize' = true -> + Int.add y z = iwordsize' -> + ror x (repr (Int.unsigned z)) = or (shl' x y) (shru' x z). +Proof. + intros. destruct (int_ltu_2_inv y z) as (A & B & C & D & E & F); auto. rewrite H1; omega. + replace (shl' x y) with (shl x (repr (Int.unsigned y))). + replace (shru' x z) with (shru x (repr (Int.unsigned z))). + apply or_ror; auto. rewrite F, H1. reflexivity. + unfold shru, shru'; rewrite <- B; auto. + unfold shl, shl'; rewrite <- A; auto. +Qed. + +Theorem shl'_shl': + forall x y z, + Int.ltu y iwordsize' = true -> + Int.ltu z iwordsize' = true -> + Int.ltu (Int.add y z) iwordsize' = true -> + shl' (shl' x y) z = shl' x (Int.add y z). +Proof. + intros. apply Int.ltu_inv in H1. + destruct (int_ltu_2_inv y z) as (A & B & C & D & E & F); auto. omega. + set (y' := repr (Int.unsigned y)) in *. + set (z' := repr (Int.unsigned z)) in *. + replace (shl' x y) with (shl x y'). + replace (shl' (shl x y') z) with (shl (shl x y') z'). + replace (shl' x (Int.add y z)) with (shl x (add y' z')). + apply shl_shl; auto. apply zlt_true. rewrite <- E. + change (unsigned iwordsize) with zwordsize. tauto. + unfold shl, shl'. rewrite E; auto. + unfold shl at 1, shl'. rewrite <- B; auto. + unfold shl, shl'; rewrite <- A; auto. +Qed. + +Theorem shru'_shru': + forall x y z, + Int.ltu y iwordsize' = true -> + Int.ltu z iwordsize' = true -> + Int.ltu (Int.add y z) iwordsize' = true -> + shru' (shru' x y) z = shru' x (Int.add y z). +Proof. + intros. apply Int.ltu_inv in H1. + destruct (int_ltu_2_inv y z) as (A & B & C & D & E & F); auto. omega. + set (y' := repr (Int.unsigned y)) in *. + set (z' := repr (Int.unsigned z)) in *. + replace (shru' x y) with (shru x y'). + replace (shru' (shru x y') z) with (shru (shru x y') z'). + replace (shru' x (Int.add y z)) with (shru x (add y' z')). + apply shru_shru; auto. apply zlt_true. rewrite <- E. + change (unsigned iwordsize) with zwordsize. tauto. + unfold shru, shru'. rewrite E; auto. + unfold shru at 1, shru'. rewrite <- B; auto. + unfold shru, shru'; rewrite <- A; auto. +Qed. + +Theorem shr'_shr': + forall x y z, + Int.ltu y iwordsize' = true -> + Int.ltu z iwordsize' = true -> + Int.ltu (Int.add y z) iwordsize' = true -> + shr' (shr' x y) z = shr' x (Int.add y z). +Proof. + intros. apply Int.ltu_inv in H1. + destruct (int_ltu_2_inv y z) as (A & B & C & D & E & F); auto. omega. + set (y' := repr (Int.unsigned y)) in *. + set (z' := repr (Int.unsigned z)) in *. + replace (shr' x y) with (shr x y'). + replace (shr' (shr x y') z) with (shr (shr x y') z'). + replace (shr' x (Int.add y z)) with (shr x (add y' z')). + apply shr_shr; auto. apply zlt_true. rewrite <- E. + change (unsigned iwordsize) with zwordsize. tauto. + unfold shr, shr'. rewrite E; auto. + unfold shr at 1, shr'. rewrite <- B; auto. + unfold shr, shr'; rewrite <- A; auto. +Qed. + +(** Powers of two with exponents given as 32-bit ints *) + +Definition one_bits' (x: int) : list Int.int := + List.map Int.repr (Z_one_bits wordsize (unsigned x) 0). + +Definition is_power2' (x: int) : option Int.int := + match Z_one_bits wordsize (unsigned x) 0 with + | i :: nil => Some (Int.repr i) + | _ => None + end. + +Theorem one_bits'_range: + forall x i, In i (one_bits' x) -> Int.ltu i iwordsize' = true. +Proof. + intros. + destruct (list_in_map_inv _ _ _ H) as [i0 [EQ IN]]. + exploit Z_one_bits_range; eauto. intros R. + unfold Int.ltu. rewrite EQ. rewrite Int.unsigned_repr. + change (Int.unsigned iwordsize') with zwordsize. apply zlt_true. omega. + assert (zwordsize < Int.max_unsigned) by reflexivity. omega. +Qed. + +Fixpoint int_of_one_bits' (l: list Int.int) : int := + match l with + | nil => zero + | a :: b => add (shl' one a) (int_of_one_bits' b) + end. + +Theorem one_bits'_decomp: + forall x, x = int_of_one_bits' (one_bits' x). +Proof. + assert (REC: forall l, + (forall i, In i l -> 0 <= i < zwordsize) -> + int_of_one_bits' (List.map Int.repr l) = repr (powerserie l)). + { induction l; simpl; intros. + - auto. + - rewrite IHl by eauto. apply eqm_samerepr; apply eqm_add. + + rewrite shl'_one_two_p. rewrite Int.unsigned_repr. apply eqm_sym; apply eqm_unsigned_repr. + exploit (H a). auto. assert (zwordsize < Int.max_unsigned) by reflexivity. omega. + + apply eqm_sym; apply eqm_unsigned_repr. + } + intros. rewrite <- (repr_unsigned x) at 1. unfold one_bits'. rewrite REC. + rewrite <- Z_one_bits_powerserie. auto. apply unsigned_range. + apply Z_one_bits_range. +Qed. + +Lemma is_power2'_rng: + forall n logn, + is_power2' n = Some logn -> + 0 <= Int.unsigned logn < zwordsize. +Proof. + unfold is_power2'; intros n logn P2. + destruct (Z_one_bits wordsize (unsigned n) 0) as [ | i [ | ? ?]] eqn:B; inv P2. + assert (0 <= i < zwordsize). + { apply Z_one_bits_range with (unsigned n). rewrite B; auto with coqlib. } + rewrite Int.unsigned_repr. auto. + assert (zwordsize < Int.max_unsigned) by reflexivity. + omega. +Qed. + +Theorem is_power2'_range: + forall n logn, + is_power2' n = Some logn -> Int.ltu logn iwordsize' = true. +Proof. + intros. unfold Int.ltu. change (Int.unsigned iwordsize') with zwordsize. + apply zlt_true. generalize (is_power2'_rng _ _ H). tauto. +Qed. + +Lemma is_power2'_correct: + forall n logn, + is_power2' n = Some logn -> + unsigned n = two_p (Int.unsigned logn). +Proof. + unfold is_power2'; intros. + destruct (Z_one_bits wordsize (unsigned n) 0) as [ | i [ | ? ?]] eqn:B; inv H. + rewrite (Z_one_bits_powerserie (unsigned n)) by (apply unsigned_range). + rewrite Int.unsigned_repr. rewrite B; simpl. omega. + assert (0 <= i < zwordsize). + { apply Z_one_bits_range with (unsigned n). rewrite B; auto with coqlib. } + assert (zwordsize < Int.max_unsigned) by reflexivity. + omega. +Qed. + +Theorem mul_pow2': + forall x n logn, + is_power2' n = Some logn -> + mul x n = shl' x logn. +Proof. + intros. rewrite shl'_mul. f_equal. rewrite shl'_one_two_p. + rewrite <- (repr_unsigned n). f_equal. apply is_power2'_correct; auto. +Qed. + +Theorem divu_pow2': + forall x n logn, + is_power2' n = Some logn -> + divu x n = shru' x logn. +Proof. + intros. generalize (is_power2'_correct n logn H). intro. + symmetry. unfold divu. rewrite H0. unfold shru'. rewrite Zshiftr_div_two_p. auto. + eapply is_power2'_rng; eauto. +Qed. + +(** Decomposing 64-bit ints as pairs of 32-bit ints *) + +Definition loword (n: int) : Int.int := Int.repr (unsigned n). + +Definition hiword (n: int) : Int.int := Int.repr (unsigned (shru n (repr Int.zwordsize))). + +Definition ofwords (hi lo: Int.int) : int := + or (shl (repr (Int.unsigned hi)) (repr Int.zwordsize)) (repr (Int.unsigned lo)). + +Lemma bits_loword: + forall n i, 0 <= i < Int.zwordsize -> Int.testbit (loword n) i = testbit n i. +Proof. + intros. unfold loword. rewrite Int.testbit_repr; auto. +Qed. + +Lemma bits_hiword: + forall n i, 0 <= i < Int.zwordsize -> Int.testbit (hiword n) i = testbit n (i + Int.zwordsize). +Proof. + intros. unfold hiword. rewrite Int.testbit_repr; auto. + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. + fold (testbit (shru n (repr Int.zwordsize)) i). rewrite bits_shru. + change (unsigned (repr Int.zwordsize)) with Int.zwordsize. + apply zlt_true. omega. omega. +Qed. + +Lemma bits_ofwords: + forall hi lo i, 0 <= i < zwordsize -> + testbit (ofwords hi lo) i = + if zlt i Int.zwordsize then Int.testbit lo i else Int.testbit hi (i - Int.zwordsize). +Proof. + intros. unfold ofwords. rewrite bits_or; auto. rewrite bits_shl; auto. + change (unsigned (repr Int.zwordsize)) with Int.zwordsize. + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. + destruct (zlt i Int.zwordsize). + rewrite testbit_repr; auto. + rewrite !testbit_repr; auto. + fold (Int.testbit lo i). rewrite Int.bits_above. apply orb_false_r. auto. + omega. +Qed. + +Lemma lo_ofwords: + forall hi lo, loword (ofwords hi lo) = lo. +Proof. + intros. apply Int.same_bits_eq; intros. + rewrite bits_loword; auto. rewrite bits_ofwords. apply zlt_true. omega. + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. omega. +Qed. + +Lemma hi_ofwords: + forall hi lo, hiword (ofwords hi lo) = hi. +Proof. + intros. apply Int.same_bits_eq; intros. + rewrite bits_hiword; auto. rewrite bits_ofwords. + rewrite zlt_false. f_equal. omega. omega. + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. omega. +Qed. + +Lemma ofwords_recompose: + forall n, ofwords (hiword n) (loword n) = n. +Proof. + intros. apply same_bits_eq; intros. rewrite bits_ofwords; auto. + destruct (zlt i Int.zwordsize). + apply bits_loword. omega. + rewrite bits_hiword. f_equal. omega. + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. omega. +Qed. + +Lemma ofwords_add: + forall lo hi, ofwords hi lo = repr (Int.unsigned hi * two_p 32 + Int.unsigned lo). +Proof. + intros. unfold ofwords. rewrite shifted_or_is_add. + apply eqm_samerepr. apply eqm_add. apply eqm_mult. + apply eqm_sym; apply eqm_unsigned_repr. + apply eqm_refl. + apply eqm_sym; apply eqm_unsigned_repr. + change Int.zwordsize with 32; change zwordsize with 64; omega. + rewrite unsigned_repr. generalize (Int.unsigned_range lo). intros [A B]. exact B. + assert (Int.max_unsigned < max_unsigned) by (compute; auto). + generalize (Int.unsigned_range_2 lo); omega. +Qed. + +Lemma ofwords_add': + forall lo hi, unsigned (ofwords hi lo) = Int.unsigned hi * two_p 32 + Int.unsigned lo. +Proof. + intros. rewrite ofwords_add. apply unsigned_repr. + generalize (Int.unsigned_range hi) (Int.unsigned_range lo). + change (two_p 32) with Int.modulus. + change Int.modulus with 4294967296. + change max_unsigned with 18446744073709551615. + omega. +Qed. + +Remark eqm_mul_2p32: + forall x y, Int.eqm x y -> eqm (x * two_p 32) (y * two_p 32). +Proof. + intros. destruct H as [k EQ]. exists k. rewrite EQ. + change Int.modulus with (two_p 32). + change modulus with (two_p 32 * two_p 32). + ring. +Qed. + +Lemma ofwords_add'': + forall lo hi, signed (ofwords hi lo) = Int.signed hi * two_p 32 + Int.unsigned lo. +Proof. + intros. rewrite ofwords_add. + replace (repr (Int.unsigned hi * two_p 32 + Int.unsigned lo)) + with (repr (Int.signed hi * two_p 32 + Int.unsigned lo)). + apply signed_repr. + generalize (Int.signed_range hi) (Int.unsigned_range lo). + change (two_p 32) with Int.modulus. + change min_signed with (Int.min_signed * Int.modulus). + change max_signed with (Int.max_signed * Int.modulus + Int.modulus - 1). + change Int.modulus with 4294967296. + omega. + apply eqm_samerepr. apply eqm_add. apply eqm_mul_2p32. apply Int.eqm_signed_unsigned. apply eqm_refl. +Qed. + +(** Expressing 64-bit operations in terms of 32-bit operations *) + +Lemma decompose_bitwise_binop: + forall f f64 f32 xh xl yh yl, + (forall x y i, 0 <= i < zwordsize -> testbit (f64 x y) i = f (testbit x i) (testbit y i)) -> + (forall x y i, 0 <= i < Int.zwordsize -> Int.testbit (f32 x y) i = f (Int.testbit x i) (Int.testbit y i)) -> + f64 (ofwords xh xl) (ofwords yh yl) = ofwords (f32 xh yh) (f32 xl yl). +Proof. + intros. apply Int64.same_bits_eq; intros. + rewrite H by auto. rewrite ! bits_ofwords by auto. + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. + destruct (zlt i Int.zwordsize); rewrite H0 by omega; auto. +Qed. + +Lemma decompose_and: + forall xh xl yh yl, + and (ofwords xh xl) (ofwords yh yl) = ofwords (Int.and xh yh) (Int.and xl yl). +Proof. + intros. apply decompose_bitwise_binop with andb. + apply bits_and. apply Int.bits_and. +Qed. + +Lemma decompose_or: + forall xh xl yh yl, + or (ofwords xh xl) (ofwords yh yl) = ofwords (Int.or xh yh) (Int.or xl yl). +Proof. + intros. apply decompose_bitwise_binop with orb. + apply bits_or. apply Int.bits_or. +Qed. + +Lemma decompose_xor: + forall xh xl yh yl, + xor (ofwords xh xl) (ofwords yh yl) = ofwords (Int.xor xh yh) (Int.xor xl yl). +Proof. + intros. apply decompose_bitwise_binop with xorb. + apply bits_xor. apply Int.bits_xor. +Qed. + +Lemma decompose_not: + forall xh xl, + not (ofwords xh xl) = ofwords (Int.not xh) (Int.not xl). +Proof. + intros. unfold not, Int.not. rewrite <- decompose_xor. f_equal. + apply (Int64.eq_spec mone (ofwords Int.mone Int.mone)). +Qed. + +Lemma decompose_shl_1: + forall xh xl y, + 0 <= Int.unsigned y < Int.zwordsize -> + shl' (ofwords xh xl) y = + ofwords (Int.or (Int.shl xh y) (Int.shru xl (Int.sub Int.iwordsize y))) + (Int.shl xl y). +Proof. + intros. + assert (Int.unsigned (Int.sub Int.iwordsize y) = Int.zwordsize - Int.unsigned y). + { unfold Int.sub. rewrite Int.unsigned_repr. auto. + rewrite Int.unsigned_repr_wordsize. generalize Int.wordsize_max_unsigned; omega. } + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. + apply Int64.same_bits_eq; intros. + rewrite bits_shl' by auto. symmetry. rewrite bits_ofwords by auto. + destruct (zlt i Int.zwordsize). rewrite Int.bits_shl by omega. + destruct (zlt i (Int.unsigned y)). auto. + rewrite bits_ofwords by omega. rewrite zlt_true by omega. auto. + rewrite zlt_false by omega. rewrite bits_ofwords by omega. + rewrite Int.bits_or by omega. rewrite Int.bits_shl by omega. + rewrite Int.bits_shru by omega. rewrite H0. + destruct (zlt (i - Int.unsigned y) (Int.zwordsize)). + rewrite zlt_true by omega. rewrite zlt_true by omega. + rewrite orb_false_l. f_equal. omega. + rewrite zlt_false by omega. rewrite zlt_false by omega. + rewrite orb_false_r. f_equal. omega. +Qed. + +Lemma decompose_shl_2: + forall xh xl y, + Int.zwordsize <= Int.unsigned y < zwordsize -> + shl' (ofwords xh xl) y = + ofwords (Int.shl xl (Int.sub y Int.iwordsize)) Int.zero. +Proof. + intros. + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. + assert (Int.unsigned (Int.sub y Int.iwordsize) = Int.unsigned y - Int.zwordsize). + { unfold Int.sub. rewrite Int.unsigned_repr. auto. + rewrite Int.unsigned_repr_wordsize. generalize (Int.unsigned_range_2 y). omega. } + apply Int64.same_bits_eq; intros. + rewrite bits_shl' by auto. symmetry. rewrite bits_ofwords by auto. + destruct (zlt i Int.zwordsize). rewrite zlt_true by omega. apply Int.bits_zero. + rewrite Int.bits_shl by omega. + destruct (zlt i (Int.unsigned y)). + rewrite zlt_true by omega. auto. + rewrite zlt_false by omega. + rewrite bits_ofwords by omega. rewrite zlt_true by omega. f_equal. omega. +Qed. + +Lemma decompose_shru_1: + forall xh xl y, + 0 <= Int.unsigned y < Int.zwordsize -> + shru' (ofwords xh xl) y = + ofwords (Int.shru xh y) + (Int.or (Int.shru xl y) (Int.shl xh (Int.sub Int.iwordsize y))). +Proof. + intros. + assert (Int.unsigned (Int.sub Int.iwordsize y) = Int.zwordsize - Int.unsigned y). + { unfold Int.sub. rewrite Int.unsigned_repr. auto. + rewrite Int.unsigned_repr_wordsize. generalize Int.wordsize_max_unsigned; omega. } + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. + apply Int64.same_bits_eq; intros. + rewrite bits_shru' by auto. symmetry. rewrite bits_ofwords by auto. + destruct (zlt i Int.zwordsize). + rewrite zlt_true by omega. + rewrite bits_ofwords by omega. + rewrite Int.bits_or by omega. rewrite Int.bits_shl by omega. + rewrite Int.bits_shru by omega. rewrite H0. + destruct (zlt (i + Int.unsigned y) (Int.zwordsize)). + rewrite zlt_true by omega. + rewrite orb_false_r. auto. + rewrite zlt_false by omega. + rewrite orb_false_l. f_equal. omega. + rewrite Int.bits_shru by omega. + destruct (zlt (i + Int.unsigned y) zwordsize). + rewrite bits_ofwords by omega. + rewrite zlt_true by omega. rewrite zlt_false by omega. f_equal. omega. + rewrite zlt_false by omega. auto. +Qed. + +Lemma decompose_shru_2: + forall xh xl y, + Int.zwordsize <= Int.unsigned y < zwordsize -> + shru' (ofwords xh xl) y = + ofwords Int.zero (Int.shru xh (Int.sub y Int.iwordsize)). +Proof. + intros. + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. + assert (Int.unsigned (Int.sub y Int.iwordsize) = Int.unsigned y - Int.zwordsize). + { unfold Int.sub. rewrite Int.unsigned_repr. auto. + rewrite Int.unsigned_repr_wordsize. generalize (Int.unsigned_range_2 y). omega. } + apply Int64.same_bits_eq; intros. + rewrite bits_shru' by auto. symmetry. rewrite bits_ofwords by auto. + destruct (zlt i Int.zwordsize). + rewrite Int.bits_shru by omega. rewrite H1. + destruct (zlt (i + Int.unsigned y) zwordsize). + rewrite zlt_true by omega. rewrite bits_ofwords by omega. + rewrite zlt_false by omega. f_equal; omega. + rewrite zlt_false by omega. auto. + rewrite zlt_false by omega. apply Int.bits_zero. +Qed. + +Lemma decompose_shr_1: + forall xh xl y, + 0 <= Int.unsigned y < Int.zwordsize -> + shr' (ofwords xh xl) y = + ofwords (Int.shr xh y) + (Int.or (Int.shru xl y) (Int.shl xh (Int.sub Int.iwordsize y))). +Proof. + intros. + assert (Int.unsigned (Int.sub Int.iwordsize y) = Int.zwordsize - Int.unsigned y). + { unfold Int.sub. rewrite Int.unsigned_repr. auto. + rewrite Int.unsigned_repr_wordsize. generalize Int.wordsize_max_unsigned; omega. } + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. + apply Int64.same_bits_eq; intros. + rewrite bits_shr' by auto. symmetry. rewrite bits_ofwords by auto. + destruct (zlt i Int.zwordsize). + rewrite zlt_true by omega. + rewrite bits_ofwords by omega. + rewrite Int.bits_or by omega. rewrite Int.bits_shl by omega. + rewrite Int.bits_shru by omega. rewrite H0. + destruct (zlt (i + Int.unsigned y) (Int.zwordsize)). + rewrite zlt_true by omega. + rewrite orb_false_r. auto. + rewrite zlt_false by omega. + rewrite orb_false_l. f_equal. omega. + rewrite Int.bits_shr by omega. + destruct (zlt (i + Int.unsigned y) zwordsize). + rewrite bits_ofwords by omega. + rewrite zlt_true by omega. rewrite zlt_false by omega. f_equal. omega. + rewrite zlt_false by omega. rewrite bits_ofwords by omega. + rewrite zlt_false by omega. f_equal. +Qed. + +Lemma decompose_shr_2: + forall xh xl y, + Int.zwordsize <= Int.unsigned y < zwordsize -> + shr' (ofwords xh xl) y = + ofwords (Int.shr xh (Int.sub Int.iwordsize Int.one)) + (Int.shr xh (Int.sub y Int.iwordsize)). +Proof. + intros. + assert (zwordsize = 2 * Int.zwordsize) by reflexivity. + assert (Int.unsigned (Int.sub y Int.iwordsize) = Int.unsigned y - Int.zwordsize). + { unfold Int.sub. rewrite Int.unsigned_repr. auto. + rewrite Int.unsigned_repr_wordsize. generalize (Int.unsigned_range_2 y). omega. } + apply Int64.same_bits_eq; intros. + rewrite bits_shr' by auto. symmetry. rewrite bits_ofwords by auto. + destruct (zlt i Int.zwordsize). + rewrite Int.bits_shr by omega. rewrite H1. + destruct (zlt (i + Int.unsigned y) zwordsize). + rewrite zlt_true by omega. rewrite bits_ofwords by omega. + rewrite zlt_false by omega. f_equal; omega. + rewrite zlt_false by omega. rewrite bits_ofwords by omega. + rewrite zlt_false by omega. auto. + rewrite Int.bits_shr by omega. + change (Int.unsigned (Int.sub Int.iwordsize Int.one)) with (Int.zwordsize - 1). + destruct (zlt (i + Int.unsigned y) zwordsize); + rewrite bits_ofwords by omega. + symmetry. rewrite zlt_false by omega. f_equal. + destruct (zlt (i - Int.zwordsize + (Int.zwordsize - 1)) Int.zwordsize); omega. + symmetry. rewrite zlt_false by omega. f_equal. + destruct (zlt (i - Int.zwordsize + (Int.zwordsize - 1)) Int.zwordsize); omega. +Qed. + +Lemma decompose_add: + forall xh xl yh yl, + add (ofwords xh xl) (ofwords yh yl) = + ofwords (Int.add (Int.add xh yh) (Int.add_carry xl yl Int.zero)) + (Int.add xl yl). +Proof. + intros. symmetry. rewrite ofwords_add. rewrite add_unsigned. + apply eqm_samerepr. + rewrite ! ofwords_add'. rewrite (Int.unsigned_add_carry xl yl). + set (cc := Int.add_carry xl yl Int.zero). + set (Xl := Int.unsigned xl); set (Xh := Int.unsigned xh); + set (Yl := Int.unsigned yl); set (Yh := Int.unsigned yh). + change Int.modulus with (two_p 32). + replace (Xh * two_p 32 + Xl + (Yh * two_p 32 + Yl)) + with ((Xh + Yh) * two_p 32 + (Xl + Yl)) by ring. + replace (Int.unsigned (Int.add (Int.add xh yh) cc) * two_p 32 + + (Xl + Yl - Int.unsigned cc * two_p 32)) + with ((Int.unsigned (Int.add (Int.add xh yh) cc) - Int.unsigned cc) * two_p 32 + + (Xl + Yl)) by ring. + apply eqm_add. 2: apply eqm_refl. apply eqm_mul_2p32. + replace (Xh + Yh) with ((Xh + Yh + Int.unsigned cc) - Int.unsigned cc) by ring. + apply Int.eqm_sub. 2: apply Int.eqm_refl. + apply Int.eqm_unsigned_repr_l. apply Int.eqm_add. 2: apply Int.eqm_refl. + apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl. +Qed. + +Lemma decompose_sub: + forall xh xl yh yl, + sub (ofwords xh xl) (ofwords yh yl) = + ofwords (Int.sub (Int.sub xh yh) (Int.sub_borrow xl yl Int.zero)) + (Int.sub xl yl). +Proof. + intros. symmetry. rewrite ofwords_add. + apply eqm_samerepr. + rewrite ! ofwords_add'. rewrite (Int.unsigned_sub_borrow xl yl). + set (bb := Int.sub_borrow xl yl Int.zero). + set (Xl := Int.unsigned xl); set (Xh := Int.unsigned xh); + set (Yl := Int.unsigned yl); set (Yh := Int.unsigned yh). + change Int.modulus with (two_p 32). + replace (Xh * two_p 32 + Xl - (Yh * two_p 32 + Yl)) + with ((Xh - Yh) * two_p 32 + (Xl - Yl)) by ring. + replace (Int.unsigned (Int.sub (Int.sub xh yh) bb) * two_p 32 + + (Xl - Yl + Int.unsigned bb * two_p 32)) + with ((Int.unsigned (Int.sub (Int.sub xh yh) bb) + Int.unsigned bb) * two_p 32 + + (Xl - Yl)) by ring. + apply eqm_add. 2: apply eqm_refl. apply eqm_mul_2p32. + replace (Xh - Yh) with ((Xh - Yh - Int.unsigned bb) + Int.unsigned bb) by ring. + apply Int.eqm_add. 2: apply Int.eqm_refl. + apply Int.eqm_unsigned_repr_l. apply Int.eqm_add. 2: apply Int.eqm_refl. + apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl. +Qed. + +Lemma decompose_sub': + forall xh xl yh yl, + sub (ofwords xh xl) (ofwords yh yl) = + ofwords (Int.add (Int.add xh (Int.not yh)) (Int.add_carry xl (Int.not yl) Int.one)) + (Int.sub xl yl). +Proof. + intros. rewrite decompose_sub. f_equal. + rewrite Int.sub_borrow_add_carry by auto. + rewrite Int.sub_add_not_3. rewrite Int.xor_assoc. rewrite Int.xor_idem. + rewrite Int.xor_zero. auto. + rewrite Int.xor_zero_l. unfold Int.add_carry. + destruct (zlt (Int.unsigned xl + Int.unsigned (Int.not yl) + Int.unsigned Int.one) Int.modulus); + compute; [right|left]; apply Int.mkint_eq; auto. +Qed. + +Definition mul' (x y: Int.int) : int := repr (Int.unsigned x * Int.unsigned y). + +Lemma mul'_mulhu: + forall x y, mul' x y = ofwords (Int.mulhu x y) (Int.mul x y). +Proof. + intros. + rewrite ofwords_add. unfold mul', Int.mulhu, Int.mul. + set (p := Int.unsigned x * Int.unsigned y). + set (ph := p / Int.modulus). set (pl := p mod Int.modulus). + transitivity (repr (ph * Int.modulus + pl)). +- f_equal. rewrite Zmult_comm. apply Z_div_mod_eq. apply Int.modulus_pos. +- apply eqm_samerepr. apply eqm_add. apply eqm_mul_2p32. auto with ints. + rewrite Int.unsigned_repr_eq. apply eqm_refl. +Qed. + +Lemma decompose_mul: + forall xh xl yh yl, + mul (ofwords xh xl) (ofwords yh yl) = + ofwords (Int.add (Int.add (hiword (mul' xl yl)) (Int.mul xl yh)) (Int.mul xh yl)) + (loword (mul' xl yl)). +Proof. + intros. + set (pl := loword (mul' xl yl)); set (ph := hiword (mul' xl yl)). + assert (EQ0: unsigned (mul' xl yl) = Int.unsigned ph * two_p 32 + Int.unsigned pl). + { rewrite <- (ofwords_recompose (mul' xl yl)). apply ofwords_add'. } + symmetry. rewrite ofwords_add. unfold mul. rewrite !ofwords_add'. + set (XL := Int.unsigned xl); set (XH := Int.unsigned xh); + set (YL := Int.unsigned yl); set (YH := Int.unsigned yh). + set (PH := Int.unsigned ph) in *. set (PL := Int.unsigned pl) in *. + transitivity (repr (((PH + XL * YH) + XH * YL) * two_p 32 + PL)). + apply eqm_samerepr. apply eqm_add. 2: apply eqm_refl. + apply eqm_mul_2p32. + rewrite Int.add_unsigned. apply Int.eqm_unsigned_repr_l. apply Int.eqm_add. + rewrite Int.add_unsigned. apply Int.eqm_unsigned_repr_l. apply Int.eqm_add. + apply Int.eqm_refl. + unfold Int.mul. apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl. + unfold Int.mul. apply Int.eqm_unsigned_repr_l. apply Int.eqm_refl. + transitivity (repr (unsigned (mul' xl yl) + (XL * YH + XH * YL) * two_p 32)). + rewrite EQ0. f_equal. ring. + transitivity (repr ((XL * YL + (XL * YH + XH * YL) * two_p 32))). + apply eqm_samerepr. apply eqm_add. 2: apply eqm_refl. + unfold mul'. apply eqm_unsigned_repr_l. apply eqm_refl. + transitivity (repr (0 + (XL * YL + (XL * YH + XH * YL) * two_p 32))). + rewrite Zplus_0_l; auto. + transitivity (repr (XH * YH * (two_p 32 * two_p 32) + (XL * YL + (XL * YH + XH * YL) * two_p 32))). + apply eqm_samerepr. apply eqm_add. 2: apply eqm_refl. + change (two_p 32 * two_p 32) with modulus. exists (- XH * YH). ring. + f_equal. ring. +Qed. + +Lemma decompose_mul_2: + forall xh xl yh yl, + mul (ofwords xh xl) (ofwords yh yl) = + ofwords (Int.add (Int.add (Int.mulhu xl yl) (Int.mul xl yh)) (Int.mul xh yl)) + (Int.mul xl yl). +Proof. + intros. rewrite decompose_mul. rewrite mul'_mulhu. + rewrite hi_ofwords, lo_ofwords. auto. +Qed. + +Lemma decompose_ltu: + forall xh xl yh yl, + ltu (ofwords xh xl) (ofwords yh yl) = if Int.eq xh yh then Int.ltu xl yl else Int.ltu xh yh. +Proof. + intros. unfold ltu. rewrite ! ofwords_add'. unfold Int.ltu, Int.eq. + destruct (zeq (Int.unsigned xh) (Int.unsigned yh)). + rewrite e. destruct (zlt (Int.unsigned xl) (Int.unsigned yl)). + apply zlt_true; omega. + apply zlt_false; omega. + change (two_p 32) with Int.modulus. + generalize (Int.unsigned_range xl) (Int.unsigned_range yl). + change Int.modulus with 4294967296. intros. + destruct (zlt (Int.unsigned xh) (Int.unsigned yh)). + apply zlt_true; omega. + apply zlt_false; omega. +Qed. + +Lemma decompose_leu: + forall xh xl yh yl, + negb (ltu (ofwords yh yl) (ofwords xh xl)) = + if Int.eq xh yh then negb (Int.ltu yl xl) else Int.ltu xh yh. +Proof. + intros. rewrite decompose_ltu. rewrite Int.eq_sym. + unfold Int.eq. destruct (zeq (Int.unsigned xh) (Int.unsigned yh)). + auto. + unfold Int.ltu. destruct (zlt (Int.unsigned xh) (Int.unsigned yh)). + rewrite zlt_false by omega; auto. + rewrite zlt_true by omega; auto. +Qed. + +Lemma decompose_lt: + forall xh xl yh yl, + lt (ofwords xh xl) (ofwords yh yl) = if Int.eq xh yh then Int.ltu xl yl else Int.lt xh yh. +Proof. + intros. unfold lt. rewrite ! ofwords_add''. rewrite Int.eq_signed. + destruct (zeq (Int.signed xh) (Int.signed yh)). + rewrite e. unfold Int.ltu. destruct (zlt (Int.unsigned xl) (Int.unsigned yl)). + apply zlt_true; omega. + apply zlt_false; omega. + change (two_p 32) with Int.modulus. + generalize (Int.unsigned_range xl) (Int.unsigned_range yl). + change Int.modulus with 4294967296. intros. + unfold Int.lt. destruct (zlt (Int.signed xh) (Int.signed yh)). + apply zlt_true; omega. + apply zlt_false; omega. +Qed. + +Lemma decompose_le: + forall xh xl yh yl, + negb (lt (ofwords yh yl) (ofwords xh xl)) = + if Int.eq xh yh then negb (Int.ltu yl xl) else Int.lt xh yh. +Proof. + intros. rewrite decompose_lt. rewrite Int.eq_sym. + rewrite Int.eq_signed. destruct (zeq (Int.signed xh) (Int.signed yh)). + auto. + unfold Int.lt. destruct (zlt (Int.signed xh) (Int.signed yh)). + rewrite zlt_false by omega; auto. + rewrite zlt_true by omega; auto. +Qed. + +(** Utility proofs for mixed 32bit and 64bit arithmetic *) + +Remark int_unsigned_range: + forall x, 0 <= Int.unsigned x <= max_unsigned. +Proof. + intros. + unfold max_unsigned. unfold modulus. + generalize (Int.unsigned_range x). + unfold Int.modulus in *. + change (wordsize) with 64%nat in *. + change (Int.wordsize) with 32%nat in *. + unfold two_power_nat. simpl. + omega. +Qed. + +Remark int_unsigned_repr: + forall x, unsigned (repr (Int.unsigned x)) = Int.unsigned x. +Proof. + intros. rewrite unsigned_repr. auto. + apply int_unsigned_range. +Qed. + +Lemma int_sub_ltu: + forall x y, + Int.ltu x y= true -> + Int.unsigned (Int.sub y x) = unsigned (sub (repr (Int.unsigned y)) (repr (Int.unsigned x))). +Proof. + intros. generalize (Int.sub_ltu x y H). intros. unfold Int.sub. unfold sub. + rewrite Int.unsigned_repr. rewrite unsigned_repr. + rewrite unsigned_repr by apply int_unsigned_range. rewrite int_unsigned_repr. reflexivity. + rewrite unsigned_repr by apply int_unsigned_range. + rewrite int_unsigned_repr. generalize (int_unsigned_range y). + omega. + generalize (Int.sub_ltu x y H). intros. + generalize (Int.unsigned_range_2 y). intros. omega. +Qed. + +End Int64. + +Strategy 0 [Wordsize_64.wordsize]. + +Notation int64 := Int64.int. + +Global Opaque Int.repr Int64.repr Byte.repr. + +(** * Specialization to offsets in pointer values *) + +(* Module Wordsize_Ptrofs. *) +(* Definition wordsize := if Archi.ptr64 then 64%nat else 32%nat. *) +(* Remark wordsize_not_zero: wordsize <> 0%nat. *) +(* Proof. unfold wordsize; destruct Archi.ptr64; congruence. Qed. *) +(* End Wordsize_Ptrofs. *) + +(* Strategy opaque [Wordsize_Ptrofs.wordsize]. *) + +(* Module Ptrofs. *) + +(* Include Make(Wordsize_Ptrofs). *) + +(* Definition to_int (x: int): Int.int := Int.repr (unsigned x). *) + +(* Definition to_int64 (x: int): Int64.int := Int64.repr (unsigned x). *) + +(* Definition of_int (x: Int.int) : int := repr (Int.unsigned x). *) + +(* Definition of_intu := of_int. *) + +(* Definition of_ints (x: Int.int) : int := repr (Int.signed x). *) + +(* Definition of_int64 (x: Int64.int) : int := repr (Int64.unsigned x). *) + +(* Definition of_int64u := of_int64. *) + +(* Definition of_int64s (x: Int64.int) : int := repr (Int64.signed x). *) + +(* Section AGREE32. *) + +(* Hypothesis _32: Archi.ptr64 = false. *) + +(* Lemma modulus_eq32: modulus = Int.modulus. *) +(* Proof. *) +(* unfold modulus, wordsize. *) +(* change Wordsize_Ptrofs.wordsize with (if Archi.ptr64 then 64%nat else 32%nat). *) +(* rewrite _32. reflexivity. *) +(* Qed. *) + +(* Lemma eqm32: *) +(* forall x y, Int.eqm x y <-> eqm x y. *) +(* Proof. *) +(* intros. unfold Int.eqm, eqm. rewrite modulus_eq32; tauto. *) +(* Qed. *) + +(* Definition agree32 (a: Ptrofs.int) (b: Int.int) : Prop := *) +(* Ptrofs.unsigned a = Int.unsigned b. *) + +(* Lemma agree32_repr: *) +(* forall i, agree32 (Ptrofs.repr i) (Int.repr i). *) +(* Proof. *) +(* intros; red. rewrite Ptrofs.unsigned_repr_eq, Int.unsigned_repr_eq. *) +(* apply f_equal2. auto. apply modulus_eq32. *) +(* Qed. *) + +(* Lemma agree32_signed: *) +(* forall a b, agree32 a b -> Ptrofs.signed a = Int.signed b. *) +(* Proof. *) +(* unfold agree32; intros. unfold signed, Int.signed, half_modulus, Int.half_modulus. *) +(* rewrite modulus_eq32. rewrite H. auto. *) +(* Qed. *) + +(* Lemma agree32_of_int: *) +(* forall b, agree32 (of_int b) b. *) +(* Proof. *) +(* unfold of_int; intros. rewrite <- (Int.repr_unsigned b) at 2. apply agree32_repr. *) +(* Qed. *) + +(* Lemma agree32_of_ints: *) +(* forall b, agree32 (of_ints b) b. *) +(* Proof. *) +(* unfold of_int; intros. rewrite <- (Int.repr_signed b) at 2. apply agree32_repr. *) +(* Qed. *) + +(* Lemma agree32_of_int_eq: *) +(* forall a b, agree32 a b -> of_int b = a. *) +(* Proof. *) +(* unfold agree32, of_int; intros. rewrite <- H. apply repr_unsigned. *) +(* Qed. *) + +(* Lemma agree32_of_ints_eq: *) +(* forall a b, agree32 a b -> of_ints b = a. *) +(* Proof. *) +(* unfold of_ints; intros. erewrite <- agree32_signed by eauto. apply repr_signed. *) +(* Qed. *) + +(* Lemma agree32_to_int: *) +(* forall a, agree32 a (to_int a). *) +(* Proof. *) +(* unfold agree32, to_int; intros. rewrite <- (agree32_repr (unsigned a)). *) +(* rewrite repr_unsigned; auto. *) +(* Qed. *) + +(* Lemma agree32_to_int_eq: *) +(* forall a b, agree32 a b -> to_int a = b. *) +(* Proof. *) +(* unfold agree32, to_int; intros. rewrite H. apply Int.repr_unsigned. *) +(* Qed. *) + +(* Lemma agree32_neg: *) +(* forall a1 b1, agree32 a1 b1 -> agree32 (Ptrofs.neg a1) (Int.neg b1). *) +(* Proof. *) +(* unfold agree32, Ptrofs.neg, Int.neg; intros. rewrite H. apply agree32_repr. *) +(* Qed. *) + +(* Lemma agree32_add: *) +(* forall a1 b1 a2 b2, *) +(* agree32 a1 b1 -> agree32 a2 b2 -> agree32 (Ptrofs.add a1 a2) (Int.add b1 b2). *) +(* Proof. *) +(* unfold agree32, Ptrofs.add, Int.add; intros. rewrite H, H0. apply agree32_repr. *) +(* Qed. *) + +(* Lemma agree32_sub: *) +(* forall a1 b1 a2 b2, *) +(* agree32 a1 b1 -> agree32 a2 b2 -> agree32 (Ptrofs.sub a1 a2) (Int.sub b1 b2). *) +(* Proof. *) +(* unfold agree32, Ptrofs.sub, Int.sub; intros. rewrite H, H0. apply agree32_repr. *) +(* Qed. *) + +(* Lemma agree32_mul: *) +(* forall a1 b1 a2 b2, *) +(* agree32 a1 b1 -> agree32 a2 b2 -> agree32 (Ptrofs.mul a1 a2) (Int.mul b1 b2). *) +(* Proof. *) +(* unfold agree32, Ptrofs.mul, Int.mul; intros. rewrite H, H0. apply agree32_repr. *) +(* Qed. *) + +(* Lemma agree32_divs: *) +(* forall a1 b1 a2 b2, *) +(* agree32 a1 b1 -> agree32 a2 b2 -> agree32 (Ptrofs.divs a1 a2) (Int.divs b1 b2). *) +(* Proof. *) +(* intros; unfold agree32, Ptrofs.divs, Int.divs. *) +(* erewrite ! agree32_signed by eauto. apply agree32_repr. *) +(* Qed. *) + +(* Lemma of_int_to_int: *) +(* forall n, of_int (to_int n) = n. *) +(* Proof. *) +(* intros; unfold of_int, to_int. apply eqm_repr_eq. rewrite <- eqm32. *) +(* apply Int.eqm_sym; apply Int.eqm_unsigned_repr. *) +(* Qed. *) + +(* End AGREE32. *) + +(* Section AGREE64. *) + +(* Hypothesis _64: Archi.ptr64 = true. *) + +(* Lemma modulus_eq64: modulus = Int64.modulus. *) +(* Proof. *) +(* unfold modulus, wordsize. *) +(* change Wordsize_Ptrofs.wordsize with (if Archi.ptr64 then 64%nat else 32%nat). *) +(* rewrite _64. reflexivity. *) +(* Qed. *) + +(* Lemma eqm64: *) +(* forall x y, Int64.eqm x y <-> eqm x y. *) +(* Proof. *) +(* intros. unfold Int64.eqm, eqm. rewrite modulus_eq64; tauto. *) +(* Qed. *) + +(* Definition agree64 (a: Ptrofs.int) (b: Int64.int) : Prop := *) +(* Ptrofs.unsigned a = Int64.unsigned b. *) + +(* Lemma agree64_repr: *) +(* forall i, agree64 (Ptrofs.repr i) (Int64.repr i). *) +(* Proof. *) +(* intros; red. rewrite Ptrofs.unsigned_repr_eq, Int64.unsigned_repr_eq. *) +(* apply f_equal2. auto. apply modulus_eq64. *) +(* Qed. *) + +(* Lemma agree64_signed: *) +(* forall a b, agree64 a b -> Ptrofs.signed a = Int64.signed b. *) +(* Proof. *) +(* unfold agree64; intros. unfold signed, Int64.signed, half_modulus, Int64.half_modulus. *) +(* rewrite modulus_eq64. rewrite H. auto. *) +(* Qed. *) + +(* Lemma agree64_of_int: *) +(* forall b, agree64 (of_int64 b) b. *) +(* Proof. *) +(* unfold of_int64; intros. rewrite <- (Int64.repr_unsigned b) at 2. apply agree64_repr. *) +(* Qed. *) + +(* Lemma agree64_of_int_eq: *) +(* forall a b, agree64 a b -> of_int64 b = a. *) +(* Proof. *) +(* unfold agree64, of_int64; intros. rewrite <- H. apply repr_unsigned. *) +(* Qed. *) + +(* Lemma agree64_to_int: *) +(* forall a, agree64 a (to_int64 a). *) +(* Proof. *) +(* unfold agree64, to_int64; intros. rewrite <- (agree64_repr (unsigned a)). *) +(* rewrite repr_unsigned; auto. *) +(* Qed. *) + +(* Lemma agree64_to_int_eq: *) +(* forall a b, agree64 a b -> to_int64 a = b. *) +(* Proof. *) +(* unfold agree64, to_int64; intros. rewrite H. apply Int64.repr_unsigned. *) +(* Qed. *) + +(* Lemma agree64_neg: *) +(* forall a1 b1, agree64 a1 b1 -> agree64 (Ptrofs.neg a1) (Int64.neg b1). *) +(* Proof. *) +(* unfold agree64, Ptrofs.neg, Int64.neg; intros. rewrite H. apply agree64_repr. *) +(* Qed. *) + +(* Lemma agree64_add: *) +(* forall a1 b1 a2 b2, *) +(* agree64 a1 b1 -> agree64 a2 b2 -> agree64 (Ptrofs.add a1 a2) (Int64.add b1 b2). *) +(* Proof. *) +(* unfold agree64, Ptrofs.add, Int.add; intros. rewrite H, H0. apply agree64_repr. *) +(* Qed. *) + +(* Lemma agree64_sub: *) +(* forall a1 b1 a2 b2, *) +(* agree64 a1 b1 -> agree64 a2 b2 -> agree64 (Ptrofs.sub a1 a2) (Int64.sub b1 b2). *) +(* Proof. *) +(* unfold agree64, Ptrofs.sub, Int.sub; intros. rewrite H, H0. apply agree64_repr. *) +(* Qed. *) + +(* Lemma agree64_mul: *) +(* forall a1 b1 a2 b2, *) +(* agree64 a1 b1 -> agree64 a2 b2 -> agree64 (Ptrofs.mul a1 a2) (Int64.mul b1 b2). *) +(* Proof. *) +(* unfold agree64, Ptrofs.mul, Int.mul; intros. rewrite H, H0. apply agree64_repr. *) +(* Qed. *) + +(* Lemma agree64_divs: *) +(* forall a1 b1 a2 b2, *) +(* agree64 a1 b1 -> agree64 a2 b2 -> agree64 (Ptrofs.divs a1 a2) (Int64.divs b1 b2). *) +(* Proof. *) +(* intros; unfold agree64, Ptrofs.divs, Int64.divs. *) +(* erewrite ! agree64_signed by eauto. apply agree64_repr. *) +(* Qed. *) + +(* Lemma of_int64_to_int64: *) +(* forall n, of_int64 (to_int64 n) = n. *) +(* Proof. *) +(* intros; unfold of_int64, to_int64. apply eqm_repr_eq. rewrite <- eqm64. *) +(* apply Int64.eqm_sym; apply Int64.eqm_unsigned_repr. *) +(* Qed. *) + +(* End AGREE64. *) + +(* Hint Resolve *) +(* agree32_repr agree32_of_int agree32_of_ints agree32_of_int_eq agree32_of_ints_eq *) +(* agree32_to_int agree32_to_int_eq agree32_neg agree32_add agree32_sub agree32_mul agree32_divs *) +(* agree64_repr agree64_of_int agree64_of_int_eq *) +(* agree64_to_int agree64_to_int_eq agree64_neg agree64_add agree64_sub agree64_mul agree64_divs : ptrofs. *) + +(* End Ptrofs. *) + +(* Strategy 0 [Wordsize_Ptrofs.wordsize]. *) + +(* Notation ptrofs := Ptrofs.int. *) + +(* Global Opaque Ptrofs.repr. *) + +Hint Resolve Int.modulus_pos Int.eqm_refl Int.eqm_refl2 Int.eqm_sym Int.eqm_trans + Int.eqm_small_eq Int.eqm_add Int.eqm_neg Int.eqm_sub Int.eqm_mult + Int.eqm_unsigned_repr Int.eqm_unsigned_repr_l Int.eqm_unsigned_repr_r + Int.unsigned_range Int.unsigned_range_2 + Int.repr_unsigned Int.repr_signed Int.unsigned_repr : ints. + +Hint Resolve Int64.modulus_pos Int64.eqm_refl Int64.eqm_refl2 Int64.eqm_sym Int64.eqm_trans + Int64.eqm_small_eq Int64.eqm_add Int64.eqm_neg Int64.eqm_sub Int64.eqm_mult + Int64.eqm_unsigned_repr Int64.eqm_unsigned_repr_l Int64.eqm_unsigned_repr_r + Int64.unsigned_range Int64.unsigned_range_2 + Int64.repr_unsigned Int64.repr_signed Int64.unsigned_repr : ints. + +(* Hint Resolve Ptrofs.modulus_pos Ptrofs.eqm_refl Ptrofs.eqm_refl2 Ptrofs.eqm_sym Ptrofs.eqm_trans *) +(* Ptrofs.eqm_small_eq Ptrofs.eqm_add Ptrofs.eqm_neg Ptrofs.eqm_sub Ptrofs.eqm_mult *) +(* Ptrofs.eqm_unsigned_repr Ptrofs.eqm_unsigned_repr_l Ptrofs.eqm_unsigned_repr_r *) +(* Ptrofs.unsigned_range Ptrofs.unsigned_range_2 *) +(* Ptrofs.repr_unsigned Ptrofs.repr_signed Ptrofs.unsigned_repr : ints. *) diff --git a/src/cclib/Maps.v b/src/cclib/Maps.v new file mode 100755 index 0000000..e4a9083 --- /dev/null +++ b/src/cclib/Maps.v @@ -0,0 +1,2889 @@ + +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Applicative finite maps are the main data structure used in this + project. A finite map associates data to keys. The two main operations + are [set k d m], which returns a map identical to [m] except that [d] + is associated to [k], and [get k m] which returns the data associated + to key [k] in map [m]. In this library, we distinguish two kinds of maps: +- Trees: the [get] operation returns an option type, either [None] + if no data is associated to the key, or [Some d] otherwise. +- Maps: the [get] operation always returns a data. If no data was explicitly + associated with the key, a default data provided at map initialization time + is returned. + + In this library, we provide efficient implementations of trees and + maps whose keys range over the type [positive] of binary positive + integers or any type that can be injected into [positive]. The + implementation is based on radix-2 search trees (uncompressed + Patricia trees) and guarantees logarithmic-time operations. An + inefficient implementation of maps as functions is also provided. +*) + +Require Import Equivalence EquivDec. +Require Import Coqlib. + +(* To avoid useless definitions of inductors in extracted code. *) +Local Unset Elimination Schemes. +Local Unset Case Analysis Schemes. + +Set Implicit Arguments. + +(** * The abstract signatures of trees *) + +Module Type TREE. + Parameter elt: Type. + Parameter elt_eq: forall (a b: elt), {a = b} + {a <> b}. + Parameter t: Type -> Type. + Parameter empty: forall (A: Type), t A. + Parameter get: forall (A: Type), elt -> t A -> option A. + Parameter set: forall (A: Type), elt -> A -> t A -> t A. + Parameter remove: forall (A: Type), elt -> t A -> t A. + + Definition get_default {A:Type} (default:A) (k:elt) (m : t A) : A := + match get k m with + | Some v => v + | NONE => default + end. + + (** The ``good variables'' properties for trees, expressing + commutations between [get], [set] and [remove]. *) + Axiom gempty: + forall (A: Type) (i: elt), get i (empty A) = None. + Axiom gss: + forall (A: Type) (i: elt) (x: A) (m: t A), get i (set i x m) = Some x. + Axiom gso: + forall (A: Type) (i j: elt) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Axiom gsspec: + forall (A: Type) (i j: elt) (x: A) (m: t A), + get i (set j x m) = if elt_eq i j then Some x else get i m. + (* We could implement the following, but it's not needed for the moment. + Hypothesis gsident: + forall (A: Type) (i: elt) (m: t A) (v: A), + get i m = Some v -> set i v m = m. + Hypothesis grident: + forall (A: Type) (i: elt) (m: t A) (v: A), + get i m = None -> remove i m = m. + *) + Axiom grs: + forall (A: Type) (i: elt) (m: t A), get i (remove i m) = None. + Axiom gro: + forall (A: Type) (i j: elt) (m: t A), + i <> j -> get i (remove j m) = get i m. + Axiom grspec: + forall (A: Type) (i j: elt) (m: t A), + get i (remove j m) = if elt_eq i j then None else get i m. + + (** Extensional equality between trees. *) + Parameter beq: forall (A: Type), (A -> A -> bool) -> t A -> t A -> bool. + Axiom beq_correct: + forall (A: Type) (eqA: A -> A -> bool) (t1 t2: t A), + beq eqA t1 t2 = true <-> + (forall (x: elt), + match get x t1, get x t2 with + | None, None => True + | Some y1, Some y2 => eqA y1 y2 = true + | _, _ => False + end). + + (** Applying a function to all data of a tree. *) + Parameter map: + forall (A B: Type), (elt -> A -> B) -> t A -> t B. + Axiom gmap: + forall (A B: Type) (f: elt -> A -> B) (i: elt) (m: t A), + get i (map f m) = option_map (f i) (get i m). + + (** Same as [map], but the function does not receive the [elt] argument. *) + Parameter map1: + forall (A B: Type), (A -> B) -> t A -> t B. + Axiom gmap1: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map1 f m) = option_map f (get i m). + + (** Applying a function pairwise to all data of two trees. *) + (* + Parameter combine: + forall (A B C: Type), (option A -> option B -> option C) -> t A -> t B -> t C. + Axiom gcombine: + forall (A B C: Type) (f: option A -> option B -> option C), + f None None = None -> + forall (m1: t A) (m2: t B) (i: elt), + get i (combine f m1 m2) = f (get i m1) (get i m2). *) + + (** Enumerating the bindings of a tree. *) + Parameter elements: + forall (A: Type), t A -> list (elt * A). + Axiom elements_correct: + forall (A: Type) (m: t A) (i: elt) (v: A), + get i m = Some v -> In (i, v) (elements m). + Axiom elements_complete: + forall (A: Type) (m: t A) (i: elt) (v: A), + In (i, v) (elements m) -> get i m = Some v. + Axiom elements_keys_norepet: + forall (A: Type) (m: t A), + list_norepet (List.map (@fst elt A) (elements m)). + Axiom elements_extensional: + forall (A: Type) (m n: t A), + (forall i, get i m = get i n) -> + elements m = elements n. + Axiom elements_remove: + forall (A: Type) i v (m: t A), + get i m = Some v -> + exists l1 l2, elements m = l1 ++ (i,v) :: l2 /\ elements (remove i m) = l1 ++ l2. + + (** Folding a function over all bindings of a tree. *) + Parameter fold: + forall (A B: Type), (B -> elt -> A -> B) -> t A -> B -> B. + Axiom fold_spec: + forall (A B: Type) (f: B -> elt -> A -> B) (v: B) (m: t A), + fold f m v = + List.fold_left (fun a p => f a (fst p) (snd p)) (elements m) v. + (** Same as [fold], but the function does not receive the [elt] argument. *) + Parameter fold1: + forall (A B: Type), (B -> A -> B) -> t A -> B -> B. + Axiom fold1_spec: + forall (A B: Type) (f: B -> A -> B) (v: B) (m: t A), + fold1 f m v = + List.fold_left (fun a p => f a (snd p)) (elements m) v. +End TREE. + +(** * The abstract signatures of maps *) + +Module Type MAP. + Parameter elt: Type. + Parameter elt_eq: forall (a b: elt), {a = b} + {a <> b}. + Parameter t: Type -> Type. + Parameter init: forall (A: Type), A -> t A. + Parameter get: forall (A: Type), elt -> t A -> A. + Parameter set: forall (A: Type), elt -> A -> t A -> t A. + Axiom gi: + forall (A: Type) (i: elt) (x: A), get i (init x) = x. + Axiom gss: + forall (A: Type) (i: elt) (x: A) (m: t A), get i (set i x m) = x. + Axiom gso: + forall (A: Type) (i j: elt) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Axiom gsspec: + forall (A: Type) (i j: elt) (x: A) (m: t A), + get i (set j x m) = if elt_eq i j then x else get i m. + Axiom gsident: + forall (A: Type) (i j: elt) (m: t A), get j (set i (get i m) m) = get j m. + Parameter map: forall (A B: Type), (A -> B) -> t A -> t B. + Axiom gmap: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map f m) = f(get i m). +End MAP. + +(** * An implementation of trees over type [positive] *) + +Module PTree <: TREE. + Definition elt := positive. + Definition elt_eq := peq. + + Inductive tree (A : Type) : Type := + | Leaf : tree A + | Node : tree A -> option A -> tree A -> tree A. + + Arguments Leaf [A]. + Arguments Node [A]. + Scheme tree_ind := Induction for tree Sort Prop. + + Definition t := tree. + + Definition empty (A : Type) := (Leaf : t A). + + Fixpoint get (A : Type) (i : positive) (m : t A) {struct i} : option A := + match m with + | Leaf => None + | Node l o r => + match i with + | xH => o + | xO ii => get ii l + | xI ii => get ii r + end + end. + + Fixpoint set (A : Type) (i : positive) (v : A) (m : t A) {struct i} : t A := + match m with + | Leaf => + match i with + | xH => Node Leaf (Some v) Leaf + | xO ii => Node (set ii v Leaf) None Leaf + | xI ii => Node Leaf None (set ii v Leaf) + end + | Node l o r => + match i with + | xH => Node l (Some v) r + | xO ii => Node (set ii v l) o r + | xI ii => Node l o (set ii v r) + end + end. + + Fixpoint remove (A : Type) (i : positive) (m : t A) {struct i} : t A := + match i with + | xH => + match m with + | Leaf => Leaf + | Node Leaf o Leaf => Leaf + | Node l o r => Node l None r + end + | xO ii => + match m with + | Leaf => Leaf + | Node l None Leaf => + match remove ii l with + | Leaf => Leaf + | mm => Node mm None Leaf + end + | Node l o r => Node (remove ii l) o r + end + | xI ii => + match m with + | Leaf => Leaf + | Node Leaf None r => + match remove ii r with + | Leaf => Leaf + | mm => Node Leaf None mm + end + | Node l o r => Node l o (remove ii r) + end + end. + + Theorem gempty: + forall (A: Type) (i: positive), get i (empty A) = None. + Proof. + induction i; simpl; auto. + Qed. + + Theorem gss: + forall (A: Type) (i: positive) (x: A) (m: t A), get i (set i x m) = Some x. + Proof. + induction i; destruct m; simpl; auto. + Qed. + + Lemma gleaf : forall (A : Type) (i : positive), get i (Leaf : t A) = None. + Proof. exact gempty. Qed. + + Theorem gso: + forall (A: Type) (i j: positive) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Proof. + induction i; intros; destruct j; destruct m; simpl; + try rewrite <- (gleaf A i); auto; try apply IHi; congruence. + Qed. + + Theorem gsspec: + forall (A: Type) (i j: positive) (x: A) (m: t A), + get i (set j x m) = if peq i j then Some x else get i m. + Proof. + intros. + destruct (peq i j); [ rewrite e; apply gss | apply gso; auto ]. + Qed. + + Theorem gsident: + forall (A: Type) (i: positive) (m: t A) (v: A), + get i m = Some v -> set i v m = m. + Proof. + induction i; intros; destruct m; simpl; simpl in H; try congruence. + rewrite (IHi m2 v H); congruence. + rewrite (IHi m1 v H); congruence. + Qed. + + Theorem set2: + forall (A: Type) (i: elt) (m: t A) (v1 v2: A), + set i v2 (set i v1 m) = set i v2 m. + Proof. + induction i; intros; destruct m; simpl; try (rewrite IHi); auto. + Qed. + + + Theorem set_swap : forall (A: Type) (i1 i2: elt) (m: t A) (v1 v2: A), + i1 <> i2 -> + set i2 v2 (set i1 v1 m) = set i1 v1 (set i2 v2 m) . + Proof. + induction i1; destruct i2; destruct m; simpl; + try reflexivity; intros; try rewrite IHi1; congruence. + Qed. + + Lemma rleaf : forall (A : Type) (i : positive), remove i (Leaf : t A) = Leaf. + Proof. destruct i; simpl; auto. Qed. + + Theorem grs: + forall (A: Type) (i: positive) (m: t A), get i (remove i m) = None. + Proof. + induction i; destruct m. + simpl; auto. + destruct m1; destruct o; destruct m2 as [ | ll oo rr]; simpl; auto. + rewrite (rleaf A i); auto. + cut (get i (remove i (Node ll oo rr)) = None). + destruct (remove i (Node ll oo rr)); auto; apply IHi. + apply IHi. + simpl; auto. + destruct m1 as [ | ll oo rr]; destruct o; destruct m2; simpl; auto. + rewrite (rleaf A i); auto. + cut (get i (remove i (Node ll oo rr)) = None). + destruct (remove i (Node ll oo rr)); auto; apply IHi. + apply IHi. + simpl; auto. + destruct m1; destruct m2; simpl; auto. + Qed. + + Theorem gro: + forall (A: Type) (i j: positive) (m: t A), + i <> j -> get i (remove j m) = get i m. + Proof. + induction i; intros; destruct j; destruct m; + try rewrite (rleaf A (xI j)); + try rewrite (rleaf A (xO j)); + try rewrite (rleaf A 1); auto; + destruct m1; destruct o; destruct m2; + simpl; + try apply IHi; try congruence; + try rewrite (rleaf A j); auto; + try rewrite (gleaf A i); auto. + cut (get i (remove j (Node m2_1 o m2_2)) = get i (Node m2_1 o m2_2)); + [ destruct (remove j (Node m2_1 o m2_2)); try rewrite (gleaf A i); auto + | apply IHi; congruence ]. + destruct (remove j (Node m1_1 o0 m1_2)); simpl; try rewrite (gleaf A i); + auto. + destruct (remove j (Node m2_1 o m2_2)); simpl; try rewrite (gleaf A i); + auto. + cut (get i (remove j (Node m1_1 o0 m1_2)) = get i (Node m1_1 o0 m1_2)); + [ destruct (remove j (Node m1_1 o0 m1_2)); try rewrite (gleaf A i); auto + | apply IHi; congruence ]. + destruct (remove j (Node m2_1 o m2_2)); simpl; try rewrite (gleaf A i); + auto. + destruct (remove j (Node m1_1 o0 m1_2)); simpl; try rewrite (gleaf A i); + auto. + Qed. + + Theorem grspec: + forall (A: Type) (i j: elt) (m: t A), + get i (remove j m) = if elt_eq i j then None else get i m. + Proof. + intros. destruct (elt_eq i j). subst j. apply grs. apply gro; auto. + Qed. + + Section BOOLEAN_EQUALITY. + + Variable A: Type. + Variable beqA: A -> A -> bool. + + Fixpoint bempty (m: t A) : bool := + match m with + | Leaf => true + | Node l None r => bempty l && bempty r + | Node l (Some _) r => false + end. + + Fixpoint beq (m1 m2: t A) {struct m1} : bool := + match m1, m2 with + | Leaf, _ => bempty m2 + | _, Leaf => bempty m1 + | Node l1 o1 r1, Node l2 o2 r2 => + match o1, o2 with + | None, None => true + | Some y1, Some y2 => beqA y1 y2 + | _, _ => false + end + && beq l1 l2 && beq r1 r2 + end. + + Lemma bempty_correct: + forall m, bempty m = true <-> (forall x, get x m = None). + Proof. + induction m; simpl. + split; intros. apply gleaf. auto. + destruct o; split; intros. + congruence. + generalize (H xH); simpl; congruence. + destruct (andb_prop _ _ H). rewrite IHm1 in H0. rewrite IHm2 in H1. + destruct x; simpl; auto. + apply andb_true_intro; split. + apply IHm1. intros; apply (H (xO x)). + apply IHm2. intros; apply (H (xI x)). + Qed. + + Lemma beq_correct: + forall m1 m2, + beq m1 m2 = true <-> + (forall (x: elt), + match get x m1, get x m2 with + | None, None => True + | Some y1, Some y2 => beqA y1 y2 = true + | _, _ => False + end). + Proof. + induction m1; intros. + - simpl. rewrite bempty_correct. split; intros. + rewrite gleaf. rewrite H. auto. + generalize (H x). rewrite gleaf. destruct (get x m2); tauto. + - destruct m2. + + unfold beq. rewrite bempty_correct. split; intros. + rewrite H. rewrite gleaf. auto. + generalize (H x). rewrite gleaf. destruct (get x (Node m1_1 o m1_2)); tauto. + + simpl. split; intros. + * destruct (andb_prop _ _ H). destruct (andb_prop _ _ H0). + rewrite IHm1_1 in H3. rewrite IHm1_2 in H1. + destruct x; simpl. apply H1. apply H3. + destruct o; destruct o0; auto || congruence. + * apply andb_true_intro. split. apply andb_true_intro. split. + generalize (H xH); simpl. destruct o; destruct o0; tauto. + apply IHm1_1. intros; apply (H (xO x)). + apply IHm1_2. intros; apply (H (xI x)). + Qed. + + End BOOLEAN_EQUALITY. + + Fixpoint prev_append (i j: positive) {struct i} : positive := + match i with + | xH => j + | xI i' => prev_append i' (xI j) + | xO i' => prev_append i' (xO j) + end. + + Definition prev (i: positive) : positive := + prev_append i xH. + + Lemma prev_append_prev i j: + prev (prev_append i j) = prev_append j i. + Proof. + revert j. unfold prev. + induction i as [i IH|i IH|]. 3: reflexivity. + intros j. simpl. rewrite IH. reflexivity. + intros j. simpl. rewrite IH. reflexivity. + Qed. + + Lemma prev_involutive i : + prev (prev i) = i. + Proof (prev_append_prev i xH). + + Lemma prev_append_inj i j j' : + prev_append i j = prev_append i j' -> j = j'. + Proof. + revert j j'. + induction i as [i Hi|i Hi|]; intros j j' H; auto; + specialize (Hi _ _ H); congruence. + Qed. + + Fixpoint xmap (A B : Type) (f : positive -> A -> B) (m : t A) (i : positive) + {struct m} : t B := + match m with + | Leaf => Leaf + | Node l o r => Node (xmap f l (xO i)) + (match o with None => None | Some x => Some (f (prev i) x) end) + (xmap f r (xI i)) + end. + + Definition map (A B : Type) (f : positive -> A -> B) m := xmap f m xH. + + Lemma xgmap: + forall (A B: Type) (f: positive -> A -> B) (i j : positive) (m: t A), + get i (xmap f m j) = option_map (f (prev (prev_append i j))) (get i m). + Proof. + induction i; intros; destruct m; simpl; auto. + Qed. + + Theorem gmap: + forall (A B: Type) (f: positive -> A -> B) (i: positive) (m: t A), + get i (map f m) = option_map (f i) (get i m). + Proof. + intros A B f i m. + unfold map. + rewrite xgmap. repeat f_equal. exact (prev_involutive i). + Qed. + + Fixpoint map1 (A B: Type) (f: A -> B) (m: t A) {struct m} : t B := + match m with + | Leaf => Leaf + | Node l o r => Node (map1 f l) (option_map f o) (map1 f r) + end. + + Theorem gmap1: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map1 f m) = option_map f (get i m). + Proof. + induction i; intros; destruct m; simpl; auto. + Qed. + + Definition Node' (A: Type) (l: t A) (x: option A) (r: t A): t A := + match l, x, r with + | Leaf, None, Leaf => Leaf + | _, _, _ => Node l x r + end. + + Lemma gnode': + forall (A: Type) (l r: t A) (x: option A) (i: positive), + get i (Node' l x r) = get i (Node l x r). + Proof. + intros. unfold Node'. + destruct l; destruct x; destruct r; auto. + destruct i; simpl; auto; rewrite gleaf; auto. + Qed. + + Fixpoint filter1 (A: Type) (pred: A -> bool) (m: t A) {struct m} : t A := + match m with + | Leaf => Leaf + | Node l o r => + let o' := match o with None => None | Some x => if pred x then o else None end in + Node' (filter1 pred l) o' (filter1 pred r) + end. + + Theorem gfilter1: + forall (A: Type) (pred: A -> bool) (i: elt) (m: t A), + get i (filter1 pred m) = + match get i m with None => None | Some x => if pred x then Some x else None end. + Proof. + intros until m. revert m i. induction m; simpl; intros. + rewrite gleaf; auto. + rewrite gnode'. destruct i; simpl; auto. destruct o; auto. + Qed. + + Section COMBINE. + + Variables A B C: Type. + Variable f: option A -> option B -> option C. + Hypothesis f_none_none: f None None = None. + + Fixpoint xcombine_l (m : t A) {struct m} : t C := + match m with + | Leaf => Leaf + | Node l o r => Node' (xcombine_l l) (f o None) (xcombine_l r) + end. + + Lemma xgcombine_l : + forall (m: t A) (i : positive), + get i (xcombine_l m) = f (get i m) None. + Proof. + induction m; intros; simpl. + repeat rewrite gleaf. auto. + rewrite gnode'. destruct i; simpl; auto. + Qed. + + Fixpoint xcombine_r (m : t B) {struct m} : t C := + match m with + | Leaf => Leaf + | Node l o r => Node' (xcombine_r l) (f None o) (xcombine_r r) + end. + + Lemma xgcombine_r : + forall (m: t B) (i : positive), + get i (xcombine_r m) = f None (get i m). + Proof. + induction m; intros; simpl. + repeat rewrite gleaf. auto. + rewrite gnode'. destruct i; simpl; auto. + Qed. + + Fixpoint combine (m1: t A) (m2: t B) {struct m1} : t C := + match m1 with + | Leaf => xcombine_r m2 + | Node l1 o1 r1 => + match m2 with + | Leaf => xcombine_l m1 + | Node l2 o2 r2 => Node' (combine l1 l2) (f o1 o2) (combine r1 r2) + end + end. + + Theorem gcombine: + forall (m1: t A) (m2: t B) (i: positive), + get i (combine m1 m2) = f (get i m1) (get i m2). + Proof. + induction m1; intros; simpl. + rewrite gleaf. apply xgcombine_r. + destruct m2; simpl. + rewrite gleaf. rewrite <- xgcombine_l. auto. + repeat rewrite gnode'. destruct i; simpl; auto. + Qed. + + End COMBINE. + + Lemma xcombine_lr : + forall (A B: Type) (f g : option A -> option A -> option B) (m : t A), + (forall (i j : option A), f i j = g j i) -> + xcombine_l f m = xcombine_r g m. + Proof. + induction m; intros; simpl; auto. + rewrite IHm1; auto. + rewrite IHm2; auto. + rewrite H; auto. + Qed. + + Theorem combine_commut: + forall (A B: Type) (f g: option A -> option A -> option B), + (forall (i j: option A), f i j = g j i) -> + forall (m1 m2: t A), + combine f m1 m2 = combine g m2 m1. + Proof. + intros A B f g EQ1. + assert (EQ2: forall (i j: option A), g i j = f j i). + intros; auto. + induction m1; intros; destruct m2; simpl; + try rewrite EQ1; + repeat rewrite (xcombine_lr f g); + repeat rewrite (xcombine_lr g f); + auto. + rewrite IHm1_1. + rewrite IHm1_2. + auto. + Qed. + + Fixpoint xelements (A : Type) (m : t A) (i : positive) + (k: list (positive * A)) {struct m} + : list (positive * A) := + match m with + | Leaf => k + | Node l None r => + xelements l (xO i) (xelements r (xI i) k) + | Node l (Some x) r => + xelements l (xO i) + ((prev i, x) :: xelements r (xI i) k) + end. + + Definition elements (A: Type) (m : t A) := xelements m xH nil. + + Remark xelements_append: + forall A (m: t A) i k1 k2, + xelements m i (k1 ++ k2) = xelements m i k1 ++ k2. + Proof. + induction m; intros; simpl. + - auto. + - destruct o; rewrite IHm2; rewrite <- IHm1; auto. + Qed. + + Remark xelements_leaf: + forall A i, xelements (@Leaf A) i nil = nil. + Proof. + intros; reflexivity. + Qed. + + Remark xelements_node: + forall A (m1: t A) o (m2: t A) i, + xelements (Node m1 o m2) i nil = + xelements m1 (xO i) nil + ++ match o with None => nil | Some v => (prev i, v) :: nil end + ++ xelements m2 (xI i) nil. + Proof. + intros. simpl. destruct o; simpl; rewrite <- xelements_append; auto. + Qed. + + Lemma xelements_incl: + forall (A: Type) (m: t A) (i : positive) k x, + In x k -> In x (xelements m i k). + Proof. + induction m; intros; simpl. + auto. + destruct o. + apply IHm1. simpl; right; auto. + auto. + Qed. + + Lemma xelements_correct: + forall (A: Type) (m: t A) (i j : positive) (v: A) k, + get i m = Some v -> In (prev (prev_append i j), v) (xelements m j k). + Proof. + induction m; intros. + rewrite (gleaf A i) in H; congruence. + destruct o; destruct i; simpl; simpl in H. + apply xelements_incl. right. auto. + auto. + inv H. apply xelements_incl. left. reflexivity. + apply xelements_incl. auto. + auto. + inv H. + Qed. + + Theorem elements_correct: + forall (A: Type) (m: t A) (i: positive) (v: A), + get i m = Some v -> In (i, v) (elements m). + Proof. + intros A m i v H. + generalize (xelements_correct m i xH nil H). rewrite prev_append_prev. exact id. + Qed. + + Lemma in_xelements: + forall (A: Type) (m: t A) (i k: positive) (v: A) , + In (k, v) (xelements m i nil) -> + exists j, k = prev (prev_append j i) /\ get j m = Some v. + Proof. + induction m; intros. + - rewrite xelements_leaf in H. contradiction. + - rewrite xelements_node in H. rewrite ! in_app_iff in H. destruct H as [P | [P | P]]. + + exploit IHm1; eauto. intros (j & Q & R). exists (xO j); auto. + + destruct o; simpl in P; intuition auto. inv H. exists xH; auto. + + exploit IHm2; eauto. intros (j & Q & R). exists (xI j); auto. + Qed. + + Theorem elements_complete: + forall (A: Type) (m: t A) (i: positive) (v: A), + In (i, v) (elements m) -> get i m = Some v. + Proof. + unfold elements. intros A m i v H. exploit in_xelements; eauto. intros (j & P & Q). + rewrite prev_append_prev in P. change i with (prev_append 1 i) in P. + exploit prev_append_inj; eauto. intros; congruence. + Qed. + + Definition xkeys (A: Type) (m: t A) (i: positive) := + List.map (@fst positive A) (xelements m i nil). + + Remark xkeys_leaf: + forall A i, xkeys (@Leaf A) i = nil. + Proof. + intros; reflexivity. + Qed. + + Remark xkeys_node: + forall A (m1: t A) o (m2: t A) i, + xkeys (Node m1 o m2) i = + xkeys m1 (xO i) + ++ match o with None => nil | Some v => prev i :: nil end + ++ xkeys m2 (xI i). + Proof. + intros. unfold xkeys. rewrite xelements_node. rewrite ! map_app. destruct o; auto. + Qed. + + Lemma in_xkeys: + forall (A: Type) (m: t A) (i k: positive), + In k (xkeys m i) -> + (exists j, k = prev (prev_append j i)). + Proof. + unfold xkeys; intros. + apply (list_in_map_inv) in H. destruct H as ((j, v) & -> & H). + exploit in_xelements; eauto. intros (k & P & Q). exists k; auto. + Qed. + + Lemma xelements_keys_norepet: + forall (A: Type) (m: t A) (i: positive), + list_norepet (xkeys m i). + Proof. + induction m; intros. + - rewrite xkeys_leaf; constructor. + - assert (NOTIN1: ~ In (prev i) (xkeys m1 (xO i))). + { red; intros. exploit in_xkeys; eauto. intros (j & EQ). + rewrite prev_append_prev in EQ. simpl in EQ. apply prev_append_inj in EQ. discriminate. } + assert (NOTIN2: ~ In (prev i) (xkeys m2 (xI i))). + { red; intros. exploit in_xkeys; eauto. intros (j & EQ). + rewrite prev_append_prev in EQ. simpl in EQ. apply prev_append_inj in EQ. discriminate. } + assert (DISJ: forall x, In x (xkeys m1 (xO i)) -> In x (xkeys m2 (xI i)) -> False). + { intros. exploit in_xkeys. eexact H. intros (j1 & EQ1). + exploit in_xkeys. eexact H0. intros (j2 & EQ2). + rewrite prev_append_prev in *. simpl in *. rewrite EQ2 in EQ1. apply prev_append_inj in EQ1. discriminate. } + rewrite xkeys_node. apply list_norepet_append. auto. + destruct o; simpl; auto. constructor; auto. + red; intros. red; intros; subst y. destruct o; simpl in H0. + destruct H0. subst x. tauto. eauto. eauto. + Qed. + + Theorem elements_keys_norepet: + forall (A: Type) (m: t A), + list_norepet (List.map (@fst elt A) (elements m)). + Proof. + intros. apply (xelements_keys_norepet m xH). + Qed. + + Remark xelements_empty: + forall (A: Type) (m: t A) i, (forall i, get i m = None) -> xelements m i nil = nil. + Proof. + induction m; intros. + auto. + rewrite xelements_node. rewrite IHm1, IHm2. destruct o; auto. + generalize (H xH); simpl; congruence. + intros. apply (H (xI i0)). + intros. apply (H (xO i0)). + Qed. + + Theorem elements_canonical_order': + forall (A B: Type) (R: A -> B -> Prop) (m: t A) (n: t B), + (forall i, option_rel R (get i m) (get i n)) -> + list_forall2 + (fun i_x i_y => fst i_x = fst i_y /\ R (snd i_x) (snd i_y)) + (elements m) (elements n). + Proof. + intros until n. unfold elements. generalize 1%positive. revert m n. + induction m; intros. + - simpl. rewrite xelements_empty. constructor. + intros. specialize (H i). rewrite gempty in H. inv H; auto. + - destruct n as [ | n1 o' n2 ]. + + rewrite (xelements_empty (Node m1 o m2)). simpl; constructor. + intros. specialize (H i). rewrite gempty in H. inv H; auto. + + rewrite ! xelements_node. repeat apply list_forall2_app. + apply IHm1. intros. apply (H (xO i)). + generalize (H xH); simpl; intros OR; inv OR. + constructor. + constructor. auto. constructor. + apply IHm2. intros. apply (H (xI i)). + Qed. + + Theorem elements_canonical_order: + forall (A B: Type) (R: A -> B -> Prop) (m: t A) (n: t B), + (forall i x, get i m = Some x -> exists y, get i n = Some y /\ R x y) -> + (forall i y, get i n = Some y -> exists x, get i m = Some x /\ R x y) -> + list_forall2 + (fun i_x i_y => fst i_x = fst i_y /\ R (snd i_x) (snd i_y)) + (elements m) (elements n). + Proof. + intros. apply elements_canonical_order'. + intros. destruct (get i m) as [x|] eqn:GM. + exploit H; eauto. intros (y & P & Q). rewrite P; constructor; auto. + destruct (get i n) as [y|] eqn:GN. + exploit H0; eauto. intros (x & P & Q). congruence. + constructor. + Qed. + + Theorem elements_extensional: + forall (A: Type) (m n: t A), + (forall i, get i m = get i n) -> + elements m = elements n. + Proof. + intros. + exploit (@elements_canonical_order' _ _ (fun (x y: A) => x = y) m n). + intros. rewrite H. destruct (get i n); constructor; auto. + induction 1. auto. destruct a1 as [a2 a3]; destruct b1 as [b2 b3]; simpl in *. + destruct H0. congruence. + Qed. + + Lemma xelements_remove: + forall (A: Type) v (m: t A) i j, + get i m = Some v -> + exists l1 l2, + xelements m j nil = l1 ++ (prev (prev_append i j), v) :: l2 + /\ xelements (remove i m) j nil = l1 ++ l2. + Proof. + induction m; intros. + - rewrite gleaf in H; discriminate. + - assert (REMOVE: xelements (remove i (Node m1 o m2)) j nil = + xelements (match i with + | xH => Node m1 None m2 + | xO ii => Node (remove ii m1) o m2 + | xI ii => Node m1 o (remove ii m2) end) + j nil). + { + destruct i; simpl remove. + destruct m1; auto. destruct o; auto. destruct (remove i m2); auto. + destruct o; auto. destruct m2; auto. destruct (remove i m1); auto. + destruct m1; auto. destruct m2; auto. + } + rewrite REMOVE. destruct i; simpl in H. + + destruct (IHm2 i (xI j) H) as (l1 & l2 & EQ & EQ'). + exists (xelements m1 (xO j) nil ++ + match o with None => nil | Some x => (prev j, x) :: nil end ++ + l1); + exists l2; split. + rewrite xelements_node, EQ, ! app_ass. auto. + rewrite xelements_node, EQ', ! app_ass. auto. + + destruct (IHm1 i (xO j) H) as (l1 & l2 & EQ & EQ'). + exists l1; + exists (l2 ++ + match o with None => nil | Some x => (prev j, x) :: nil end ++ + xelements m2 (xI j) nil); + split. + rewrite xelements_node, EQ, ! app_ass. auto. + rewrite xelements_node, EQ', ! app_ass. auto. + + subst o. exists (xelements m1 (xO j) nil); exists (xelements m2 (xI j) nil); split. + rewrite xelements_node. rewrite prev_append_prev. auto. + rewrite xelements_node; auto. + Qed. + + Theorem elements_remove: + forall (A: Type) i v (m: t A), + get i m = Some v -> + exists l1 l2, elements m = l1 ++ (i,v) :: l2 /\ elements (remove i m) = l1 ++ l2. + Proof. + intros. exploit xelements_remove. eauto. instantiate (1 := xH). + rewrite prev_append_prev. auto. + Qed. + + Fixpoint xfold (A B: Type) (f: B -> positive -> A -> B) + (i: positive) (m: t A) (v: B) {struct m} : B := + match m with + | Leaf => v + | Node l None r => + let v1 := xfold f (xO i) l v in + xfold f (xI i) r v1 + | Node l (Some x) r => + let v1 := xfold f (xO i) l v in + let v2 := f v1 (prev i) x in + xfold f (xI i) r v2 + end. + + Definition fold (A B : Type) (f: B -> positive -> A -> B) (m: t A) (v: B) := + xfold f xH m v. + + Lemma xfold_xelements: + forall (A B: Type) (f: B -> positive -> A -> B) m i v l, + List.fold_left (fun a p => f a (fst p) (snd p)) l (xfold f i m v) = + List.fold_left (fun a p => f a (fst p) (snd p)) (xelements m i l) v. + Proof. + induction m; intros. + simpl. auto. + destruct o; simpl. + rewrite <- IHm1. simpl. rewrite <- IHm2. auto. + rewrite <- IHm1. rewrite <- IHm2. auto. + Qed. + + Theorem fold_spec: + forall (A B: Type) (f: B -> positive -> A -> B) (v: B) (m: t A), + fold f m v = + List.fold_left (fun a p => f a (fst p) (snd p)) (elements m) v. + Proof. + intros. unfold fold, elements. rewrite <- xfold_xelements. auto. + Qed. + + Fixpoint fold1 (A B: Type) (f: B -> A -> B) (m: t A) (v: B) {struct m} : B := + match m with + | Leaf => v + | Node l None r => + let v1 := fold1 f l v in + fold1 f r v1 + | Node l (Some x) r => + let v1 := fold1 f l v in + let v2 := f v1 x in + fold1 f r v2 + end. + + Lemma fold1_xelements: + forall (A B: Type) (f: B -> A -> B) m i v l, + List.fold_left (fun a p => f a (snd p)) l (fold1 f m v) = + List.fold_left (fun a p => f a (snd p)) (xelements m i l) v. + Proof. + induction m; intros. + simpl. auto. + destruct o; simpl. + rewrite <- IHm1. simpl. rewrite <- IHm2. auto. + rewrite <- IHm1. rewrite <- IHm2. auto. + Qed. + + Theorem fold1_spec: + forall (A B: Type) (f: B -> A -> B) (v: B) (m: t A), + fold1 f m v = + List.fold_left (fun a p => f a (snd p)) (elements m) v. + Proof. + intros. apply fold1_xelements with (l := @nil (positive * A)). + Qed. + + (* This is used in the proofless backend as a simple way to union two globalenvs. + It will not be well-behaved if there is overlap between the keys, and we + currently don't prove any theorems about it. *) + Definition union (A:Type) (m1 m2 : t A) : t A := + fold (fun m k v => set k v m) m1 m2. + + Definition get_default {A:Type} (default:A) (k:elt) (m : t A) : A := + match get k m with + | Some v => v + | NONE => default + end. + +End PTree. + +(** * An implementation of maps over type [positive] *) + +Module PMap <: MAP. + Definition elt := positive. + Definition elt_eq := peq. + + Definition t (A : Type) : Type := (A * PTree.t A)%type. + + Definition init (A : Type) (x : A) := + (x, PTree.empty A). + + Definition get (A : Type) (i : positive) (m : t A) := + match PTree.get i (snd m) with + | Some x => x + | None => fst m + end. + + Definition set (A : Type) (i : positive) (x : A) (m : t A) := + (fst m, PTree.set i x (snd m)). + + Theorem gi: + forall (A: Type) (i: positive) (x: A), get i (init x) = x. + Proof. + intros. unfold init. unfold get. simpl. rewrite PTree.gempty. auto. + Qed. + + Theorem gss: + forall (A: Type) (i: positive) (x: A) (m: t A), get i (set i x m) = x. + Proof. + intros. unfold get. unfold set. simpl. rewrite PTree.gss. auto. + Qed. + + Theorem gso: + forall (A: Type) (i j: positive) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Proof. + intros. unfold get. unfold set. simpl. rewrite PTree.gso; auto. + Qed. + + Theorem gsspec: + forall (A: Type) (i j: positive) (x: A) (m: t A), + get i (set j x m) = if peq i j then x else get i m. + Proof. + intros. destruct (peq i j). + rewrite e. apply gss. auto. + apply gso. auto. + Qed. + + Theorem gsident: + forall (A: Type) (i j: positive) (m: t A), + get j (set i (get i m) m) = get j m. + Proof. + intros. destruct (peq i j). + rewrite e. rewrite gss. auto. + rewrite gso; auto. + Qed. + + Definition map (A B : Type) (f : A -> B) (m : t A) : t B := + (f (fst m), PTree.map1 f (snd m)). + + Theorem gmap: + forall (A B: Type) (f: A -> B) (i: positive) (m: t A), + get i (map f m) = f(get i m). + Proof. + intros. unfold map. unfold get. simpl. rewrite PTree.gmap1. + unfold option_map. destruct (PTree.get i (snd m)); auto. + Qed. + + Theorem set2: + forall (A: Type) (i: elt) (x y: A) (m: t A), + set i y (set i x m) = set i y m. + Proof. + intros. unfold set. simpl. decEq. apply PTree.set2. + Qed. + +End PMap. + +(** * An implementation of maps over any type that injects into type [positive] *) + +Module Type INDEXED_TYPE. + Parameter t: Type. + Parameter index: t -> positive. + Parameter index_inv : positive -> t. + Axiom index_invertible : forall x, index_inv (index x) = x. + Axiom index_inj: forall (x y: t), index x = index y -> x = y. + + Parameter eq: forall (x y: t), {x = y} + {x <> y}. +End INDEXED_TYPE. + +Module IMap(X: INDEXED_TYPE). + + Definition elt := X.t. + Definition elt_eq := X.eq. + Definition t : Type -> Type := PMap.t. + Definition init (A: Type) (x: A) := PMap.init x. + Definition get (A: Type) (i: X.t) (m: t A) := PMap.get (X.index i) m. + Definition set (A: Type) (i: X.t) (v: A) (m: t A) := PMap.set (X.index i) v m. + Definition map (A B: Type) (f: A -> B) (m: t A) : t B := PMap.map f m. + + Lemma gi: + forall (A: Type) (x: A) (i: X.t), get i (init x) = x. + Proof. + intros. unfold get, init. apply PMap.gi. + Qed. + + Lemma gss: + forall (A: Type) (i: X.t) (x: A) (m: t A), get i (set i x m) = x. + Proof. + intros. unfold get, set. apply PMap.gss. + Qed. + + Lemma gso: + forall (A: Type) (i j: X.t) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Proof. + intros. unfold get, set. apply PMap.gso. + red. intro. apply H. apply X.index_inj; auto. + Qed. + + Lemma gsspec: + forall (A: Type) (i j: X.t) (x: A) (m: t A), + get i (set j x m) = if X.eq i j then x else get i m. + Proof. + intros. unfold get, set. + rewrite PMap.gsspec. + case (X.eq i j); intro. + subst j. rewrite peq_true. reflexivity. + rewrite peq_false. reflexivity. + red; intro. elim n. apply X.index_inj; auto. + Qed. + + Lemma gmap: + forall (A B: Type) (f: A -> B) (i: X.t) (m: t A), + get i (map f m) = f(get i m). + Proof. + intros. unfold map, get. apply PMap.gmap. + Qed. + + Lemma set2: + forall (A: Type) (i: elt) (x y: A) (m: t A), + set i y (set i x m) = set i y m. + Proof. + intros. unfold set. apply PMap.set2. + Qed. + + (* This is used in the proofless backend as a simple way to union two globalenvs. + It will not be well-behaved if there is overlap between the keys, and we + currently don't prove any theorems about it. *) + Definition union (A:Type) (m1 m2 : t A) : t A := + (fst m1, PTree.union (snd m1) (snd m2)). +End IMap. + +Module ZIndexed. + Definition t := Z. + Definition index (z: Z): positive := + match z with + | Z0 => xH + | Zpos p => xO p + | Zneg p => xI p + end. + + Definition index_inv (p:positive) : Z := + match p with + | xH => Z0 + | xO p' => Zpos p' + | xI p' => Zneg p' + end. + + Lemma index_invertible : forall x, + index_inv (index x) = x. + Proof. + intros; destruct x; reflexivity. + Qed. + + Lemma index_inj: forall (x y: Z), index x = index y -> x = y. + Proof. + intros. + assert (H0: index_inv (index x) = index_inv (index y)) by congruence. + rewrite ?index_invertible in H0. + assumption. + Qed. + + Definition eq := zeq. +End ZIndexed. + +Module ZMap := IMap(ZIndexed). + +Module NIndexed. + Definition t := N. + Definition index (n: N): positive := + match n with + | N0 => xH + | Npos p => xO p + end. + + Definition index_inv (p:positive) : N := + match p with + | xH => N0 + | xO p' => Npos p' + | xI _ => N0 + end. + + Lemma index_invertible : forall x, + index_inv (index x) = x. + Proof. + intros; destruct x; reflexivity. + Qed. + + Lemma index_inj: forall (x y: N), index x = index y -> x = y. + Proof. + intros. + assert (H0: index_inv (index x) = index_inv (index y)) by congruence. + rewrite ?index_invertible in H0. + assumption. + Qed. + + + Lemma eq: forall (x y: N), {x = y} + {x <> y}. + Proof. + decide equality. apply peq. + Qed. +End NIndexed. + +Module NMap := IMap(NIndexed). + +(** * An implementation of maps over any type with decidable equality *) + +Module Type EQUALITY_TYPE. + Parameter t: Type. + Parameter eq: forall (x y: t), {x = y} + {x <> y}. +End EQUALITY_TYPE. + +Module EMap(X: EQUALITY_TYPE) <: MAP. + + Definition elt := X.t. + Definition elt_eq := X.eq. + Definition t (A: Type) := X.t -> A. + Definition init (A: Type) (v: A) := fun (_: X.t) => v. + Definition get (A: Type) (x: X.t) (m: t A) := m x. + Definition set (A: Type) (x: X.t) (v: A) (m: t A) := + fun (y: X.t) => if X.eq y x then v else m y. + Lemma gi: + forall (A: Type) (i: elt) (x: A), init x i = x. + Proof. + intros. reflexivity. + Qed. + Lemma gss: + forall (A: Type) (i: elt) (x: A) (m: t A), (set i x m) i = x. + Proof. + intros. unfold set. case (X.eq i i); intro. + reflexivity. tauto. + Qed. + Lemma gso: + forall (A: Type) (i j: elt) (x: A) (m: t A), + i <> j -> (set j x m) i = m i. + Proof. + intros. unfold set. case (X.eq i j); intro. + congruence. reflexivity. + Qed. + Lemma gsspec: + forall (A: Type) (i j: elt) (x: A) (m: t A), + get i (set j x m) = if elt_eq i j then x else get i m. + Proof. + intros. unfold get, set, elt_eq. reflexivity. + Qed. + Lemma gsident: + forall (A: Type) (i j: elt) (m: t A), get j (set i (get i m) m) = get j m. + Proof. + intros. unfold get, set. case (X.eq j i); intro. + congruence. reflexivity. + Qed. + Definition map (A B: Type) (f: A -> B) (m: t A) := + fun (x: X.t) => f(m x). + Lemma gmap: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map f m) = f(get i m). + Proof. + intros. unfold get, map. reflexivity. + Qed. +End EMap. + +(** * An implementation of trees over any type that injects into type [positive] *) + +Module ITree(X: INDEXED_TYPE) <: TREE. + + Definition elt := X.t. + Definition elt_eq := X.eq. + Definition t (A: Type) : Type := + { tr : PTree.t A | forall x, PTree.get x tr <> None -> + X.index (X.index_inv x) = x }. + + Program Definition empty (A: Type): t A := PTree.empty A. + Next Obligation. + rewrite PTree.gempty in H. + congruence. + Qed. + + Program Definition get (A: Type) (k: elt) (m: t A): option A := PTree.get (X.index k) m. + + Definition get_default {A:Type} (default:A) (k:elt) (m : t A) : A := + match get k m with + | Some v => v + | NONE => default + end. + + Program Definition set (A: Type) (k: elt) (v: A) (m: t A): t A + := PTree.set (X.index k) v m. + Next Obligation. + destruct m as [m INV]; simpl in H. + destruct (peq x (X.index k)). + + rewrite e. + rewrite X.index_invertible. + reflexivity. + + rewrite PTree.gso in H by assumption. + auto. + Qed. + + Program Definition remove (A: Type) (k: elt) (m: t A): t A + := PTree.remove (X.index k) m. + Next Obligation. + destruct m as [m INV]; simpl in H. + destruct (peq x (X.index k)). + + rewrite e. + rewrite X.index_invertible. + reflexivity. + + rewrite PTree.gro in H by assumption. + auto. + Qed. + + Theorem gempty: + forall (A: Type) (i: elt), get i (empty A) = None. + Proof. + intros. apply PTree.gempty. + Qed. + + Theorem gss: + forall (A: Type) (i: elt) (x: A) (m: t A), get i (set i x m) = Some x. + Proof. + intros. apply PTree.gss. + Qed. + + Theorem gso: + forall (A: Type) (i j: elt) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Proof. + intros. apply PTree.gso. red; intros; elim H; apply X.index_inj; auto. + Qed. + + Theorem gsspec: + forall (A: Type) (i j: elt) (x: A) (m: t A), + get i (set j x m) = if elt_eq i j then Some x else get i m. + Proof. + intros. destruct (elt_eq i j). subst j; apply gss. apply gso; auto. + Qed. + + Theorem grs: + forall (A: Type) (i: elt) (m: t A), get i (remove i m) = None. + Proof. + intros. apply PTree.grs. + Qed. + + Theorem gro: + forall (A: Type) (i j: elt) (m: t A), + i <> j -> get i (remove j m) = get i m. + Proof. + intros. apply PTree.gro. red; intros; elim H; apply X.index_inj; auto. + Qed. + + Theorem grspec: + forall (A: Type) (i j: elt) (m: t A), + get i (remove j m) = if elt_eq i j then None else get i m. + Proof. + intros. destruct (elt_eq i j). subst j; apply grs. apply gro; auto. + Qed. + + Program Definition beq: forall (A: Type), (A -> A -> bool) -> t A -> t A -> bool + := PTree.beq. + + Theorem beq_correct: + forall (A: Type) (eqA: A -> A -> bool) (t1 t2: t A), + beq eqA t1 t2 = true <-> + forall (x: elt), + match get x t1, get x t2 with + | None, None => True + | Some y1, Some y2 => eqA y1 y2 = true + | _, _ => False + end. + Proof. + unfold beq, get. split; intros. + + rewrite PTree.beq_correct in H. apply H. + + rewrite PTree.beq_correct. + destruct t1 as [t1 INV1]. + destruct t2 as [t2 INV2]. + simpl in *. + intros x. + specialize (H (X.index_inv x)). + destruct (PTree.get x t1) eqn:e1; destruct (PTree.get x t2) eqn:e2; simpl. + * rewrite (INV1 x), e1, e2 in H; congruence. + * rewrite (INV1 x), e1, e2 in H; congruence. + * rewrite (INV2 x), e1, e2 in H; congruence. + * auto. + Qed. + + + (* This probably could be defined, but it seems annoying...*) + (* + Program Definition combine: forall (A B C: Type), (option A -> option B -> option C) -> t A -> t B -> t C := PTree.combine. + Next Obligation. + destruct x3 as [m1 INV1]. + destruct x4 as [m2 INV2]. + simpl in *. + + + Theorem gcombine: + forall (A B C: Type) (f: option A -> option B -> option C), + f None None = None -> + forall (m1: t A) (m2: t B) (i: elt), + get i (combine f m1 m2) = f (get i m1) (get i m2). + Proof. + intros. apply PTree.gcombine. auto. + Qed. + *) + + Program Definition map (A B: Type) (f : elt -> A -> B) (m: t A) : t B := + PTree.map (fun p a => f (X.index_inv p) a) m. + Next Obligation. + destruct m as [m INV]. + simpl in H. + apply INV. + rewrite PTree.gmap in H. + intros C. + rewrite C in H. + simpl in H. + congruence. + Qed. + + Theorem gmap: + forall (A B: Type) (f: elt -> A -> B) (i: elt) (m: t A), + get i (map f m) = option_map (f i) (get i m). + Proof. + intros. + unfold get, map. + replace (f i) with (f (X.index_inv (X.index i))) + by (f_equal; exact (X.index_invertible i)). + change (f (X.index_inv (X.index i))) + with ((fun (p : positive) (a : A) => f (X.index_inv p) a) (X.index i)). + apply PTree.gmap. + Qed. + + Program Definition map1: + forall (A B: Type), (A -> B) -> t A -> t B := PTree.map1. + Next Obligation. + destruct x2 as [m INV]. + simpl in H. + apply INV. + rewrite PTree.gmap1 in H. + intros C. + rewrite C in H. + simpl in H. + congruence. + Qed. + + Theorem gmap1: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map1 f m) = option_map f (get i m). + Proof. + unfold get, map1. + intros. + apply PTree.gmap1. + Qed. + + (** Enumerating the bindings of a tree. *) + Program Definition elements (A: Type) (m: t A) : list (elt * A) + := List.map (fun xA => (X.index_inv (fst xA), snd xA)) (PTree.elements m). + + Theorem elements_correct: + forall (A: Type) (m: t A) (i: elt) (v: A), + get i m = Some v -> In (i, v) (elements m). + Proof. + unfold get, elements. + intros. + apply PTree.elements_correct in H. + rewrite List.in_map_iff. + exists (X.index i, v). + split; auto. + simpl. + rewrite X.index_invertible. + reflexivity. + Qed. + + Theorem elements_complete: + forall (A: Type) (m: t A) (i: elt) (v: A), + In (i, v) (elements m) -> get i m = Some v. + Proof. + unfold elements, get. + intros. + destruct m as [m INV]. + simpl in *. + apply PTree.elements_complete. + rewrite in_map_iff in H. + destruct H as [[p u] [H1 H2]]. + simpl in H1. + inversion H1. subst. + rewrite INV. + + assumption. + + apply PTree.elements_complete in H2. + congruence. + Qed. + + Theorem elements_keys_norepet: + forall (A: Type) (m: t A), + list_norepet (List.map (@fst elt A) (elements m)). + Proof. + intros. + destruct m as [m INV]. + unfold elements. + rewrite map_map. + simpl. + rewrite <- (@map_map _ _ _ fst). + apply list_map_norepet. + + apply PTree.elements_keys_norepet. + + intros. + intros C. + assert (X.index (X.index_inv x) = X.index (X.index_inv y)) by congruence. + rewrite (INV x) in H2. + rewrite (INV y) in H2. + * congruence. + * rewrite in_map_iff in H0. + destruct H0 as [[k a] [Hfst Hin]]. + apply PTree.elements_complete in Hin. + simpl in Hfst; inversion Hfst; subst. + congruence. + * rewrite in_map_iff in H. + destruct H as [[k a] [Hfst Hin]]. + apply PTree.elements_complete in Hin. + simpl in Hfst; inversion Hfst; subst. + congruence. + Qed. + + Theorem elements_extensional: + forall (A: Type) (m n: t A), + (forall i, get i m = get i n) -> + elements m = elements n. + Proof. + unfold get, elements. + destruct m as [m INVm]. + destruct n as [n INVn]. + simpl. + intros. + rewrite (PTree.elements_extensional n m); [reflexivity|]. + intros i. + specialize (H (X.index_inv i)). + destruct (PTree.get i n) eqn:?; + destruct (PTree.get i m) eqn:?. + * rewrite (INVm i) in H; congruence. + * rewrite (INVn i) in H; congruence. + * rewrite (INVm i) in H; congruence. + * reflexivity. + Qed. + + Theorem elements_remove: + forall (A: Type) i v (m: t A), + get i m = Some v -> + exists l1 l2, elements m = l1 ++ (i,v) :: l2 /\ elements (remove i m) = l1 ++ l2. + Proof. + unfold get, remove, elements. + intros. + destruct m as [m INV]. + simpl in *. + assert (H' := PTree.elements_remove _ _ H). + destruct H' as [l1 [l2 [H1 H2]]]. + exists (List.map (fun xA : positive * A => (X.index_inv (fst xA), snd xA)) l1). + exists (List.map (fun xA : positive * A => (X.index_inv (fst xA), snd xA)) l2). + rewrite H1. + rewrite H2. + repeat rewrite map_app. + simpl. + rewrite X.index_invertible. + split; reflexivity. + Qed. + + (** Folding a function over all bindings of a tree. *) + Program Definition fold (A B: Type) (f : B -> elt -> A -> B) (m : t A) (b: B) : B + := PTree.fold (fun b x a => f b (X.index_inv x) a) m b. + + Lemma fold_left_equal : forall A B (f g : A->B->A) l v, + (forall x y, f x y = g x y) -> + fold_left f l v = fold_left g l v. + Proof. + induction l. + - reflexivity. + - simpl. intros. + rewrite (IHl _ H). + rewrite H. + reflexivity. + Qed. + + Lemma fold_left_map : forall A B C (f : A->B->A) (g : C -> B) (l : list C) v, + fold_left f (List.map g l) v = fold_left (fun x y => f x (g y)) l v. + Proof. + induction l. + - reflexivity. + - simpl. intros. + rewrite IHl. + reflexivity. + Qed. + + Theorem fold_spec: + forall (A B: Type) (f: B -> elt -> A -> B) (v: B) (m: t A), + fold f m v = + List.fold_left (fun a p => f a (fst p) (snd p)) (elements m) v. + Proof. + intros. + unfold fold, elements. + rewrite PTree.fold_spec. + rewrite fold_left_map. + apply fold_left_equal. + reflexivity. + Qed. + + Program Definition fold1 (A B: Type) (f : B -> A -> B) (m: t A) (v: B) : B + := PTree.fold1 f m v. + + Theorem fold1_spec: + forall (A B: Type) (f: B -> A -> B) (v: B) (m: t A), + fold1 f m v = + List.fold_left (fun a p => f a (snd p)) (elements m) v. + Proof. + intros. + unfold fold1, elements. + rewrite PTree.fold1_spec. + rewrite fold_left_map. + apply fold_left_equal. + reflexivity. + Qed. + +End ITree. + +Module ZTree := ITree(ZIndexed). + +Require Import cclib.Integers. + +Module Int256Indexed <: INDEXED_TYPE. + +Definition t := int256. +Definition index (i : int256) := + let (intval, _) := i in + match intval with + | Z0 => xH + | Zpos p => xO p + | Zneg p => xI p + end. + +Definition index_inv p := + match p with + | xH => Int256.zero + | xO p' => Int256.repr (Zpos p') + | xI p' => Int256.repr (Zneg p') + end. + + Lemma zlt_proof_irrelevance : + forall {x y : Z} (p q : x < y), p = q. + Proof. + intros. + apply Eqdep_dec.eq_proofs_unicity. + decide equality. + Qed. + + Lemma range_proof_irrelevance : + forall {x y z : Z} (p q : x < y < z), p = q. + Proof. + intros. + destruct p as [p1 p2]. + destruct q as [q1 q2]. + rewrite (zlt_proof_irrelevance p1 q1). + rewrite (zlt_proof_irrelevance p2 q2). + reflexivity. + Qed. + + Transparent Int256.Z_mod_modulus. + Lemma modulus_nop : forall x, + -1 < x < Int256.modulus -> Int256.Z_mod_modulus x = x. + Proof. + intros. + unfold Int256.Z_mod_modulus. + destruct x. + + reflexivity. + + rewrite Int256.P_mod_two_p_eq. + rewrite Zmod_small by (unfold Int256.modulus in H; omega). + reflexivity. + + destruct H. + unfold Z.lt in H. + simpl in H. + destruct p; simpl in H; congruence. + Qed. + Opaque Int256.Z_mod_modulus. + + Transparent Int256.repr. + Lemma index_invertible : forall x, index_inv (index x) = x. + Proof. + intros. + destruct x as [x x_range]. + destruct x; simpl. + + unfold Int256.zero. + unfold Int256.repr. + rewrite (range_proof_irrelevance (Int256.Z_mod_modulus_range' 0) x_range). + reflexivity. + + unfold Int256.repr. + generalize (Int256.Z_mod_modulus_range' (Z.pos p)). + rewrite (modulus_nop x_range). + intros a. + rewrite (range_proof_irrelevance a x_range). + reflexivity. + + unfold Int256.repr. + generalize (Int256.Z_mod_modulus_range' (Z.neg p)). + rewrite (modulus_nop x_range). + intros a. + rewrite (range_proof_irrelevance a x_range). + reflexivity. + Qed. + Opaque Int256.repr. + + Lemma index_inj: forall (x y: int256), index x = index y -> x = y. + Proof. + unfold index; destruct x; destruct y; intros; + try discriminate; try reflexivity. + apply Int256.mkint_eq. + destruct intval; destruct intval0; try (inversion H); try reflexivity. + Qed. + + Definition eq := Int256.eq_dec. + +End Int256Indexed. + +Module Int256Tree := ITree(Int256Indexed). +Module Int256Map := IMap(Int256Indexed). + +(** * Additional properties over trees *) + +Module Tree_Properties(T: TREE). + + (** An induction principle over [fold]. *) + + Section TREE_FOLD_IND. + + Variables V A: Type. + Variable f: A -> T.elt -> V -> A. + Variable P: T.t V -> A -> Prop. + Variable init: A. + Variable m_final: T.t V. + + Hypothesis P_compat: + forall m m' a, + (forall x, T.get x m = T.get x m') -> + P m a -> P m' a. + + Hypothesis H_base: + P (T.empty _) init. + + Hypothesis H_rec: + forall m a k v, + T.get k m = None -> T.get k m_final = Some v -> P m a -> P (T.set k v m) (f a k v). + + Let f' (a: A) (p : T.elt * V) := f a (fst p) (snd p). + + Let P' (l: list (T.elt * V)) (a: A) : Prop := + forall m, list_equiv l (T.elements m) -> P m a. + + Remark H_base': + P' nil init. + Proof. + red; intros. apply P_compat with (T.empty _); auto. + intros. rewrite T.gempty. symmetry. case_eq (T.get x m); intros; auto. + assert (In (x, v) nil). rewrite (H (x, v)). apply T.elements_correct. auto. + contradiction. + Qed. + + Remark H_rec': + forall k v l a, + ~In k (List.map (@fst T.elt V) l) -> + In (k, v) (T.elements m_final) -> + P' l a -> + P' (l ++ (k, v) :: nil) (f a k v). + Proof. + unfold P'; intros. + set (m0 := T.remove k m). + apply P_compat with (T.set k v m0). + intros. unfold m0. rewrite T.gsspec. destruct (T.elt_eq x k). + symmetry. apply T.elements_complete. rewrite <- (H2 (x, v)). + apply in_or_app. simpl. intuition congruence. + apply T.gro. auto. + apply H_rec. unfold m0. apply T.grs. apply T.elements_complete. auto. + apply H1. red. intros [k' v']. + split; intros. + apply T.elements_correct. unfold m0. rewrite T.gro. apply T.elements_complete. + rewrite <- (H2 (k', v')). apply in_or_app. auto. + red; intro; subst k'. elim H. change k with (fst (k, v')). apply in_map. auto. + assert (T.get k' m0 = Some v'). apply T.elements_complete. auto. + unfold m0 in H4. rewrite T.grspec in H4. destruct (T.elt_eq k' k). congruence. + assert (In (k', v') (T.elements m)). apply T.elements_correct; auto. + rewrite <- (H2 (k', v')) in H5. destruct (in_app_or _ _ _ H5). auto. + simpl in H6. intuition congruence. + Qed. + + Lemma fold_rec_aux: + forall l1 l2 a, + list_equiv (l2 ++ l1) (T.elements m_final) -> + list_disjoint (List.map (@fst T.elt V) l1) (List.map (@fst T.elt V) l2) -> + list_norepet (List.map (@fst T.elt V) l1) -> + P' l2 a -> P' (l2 ++ l1) (List.fold_left f' l1 a). + Proof. + induction l1; intros; simpl. + rewrite <- List.app_nil_end. auto. + destruct a as [k v]; simpl in *. inv H1. + change ((k, v) :: l1) with (((k, v) :: nil) ++ l1). rewrite <- List.app_ass. apply IHl1. + rewrite app_ass. auto. + red; intros. rewrite map_app in H3. destruct (in_app_or _ _ _ H3). apply H0; auto with coqlib. + simpl in H4. intuition congruence. + auto. + unfold f'. simpl. apply H_rec'; auto. eapply list_disjoint_notin; eauto with coqlib. + rewrite <- (H (k, v)). apply in_or_app. simpl. auto. + Qed. + + Theorem fold_rec: + P m_final (T.fold f m_final init). + Proof. + intros. rewrite T.fold_spec. fold f'. + assert (P' (nil ++ T.elements m_final) (List.fold_left f' (T.elements m_final) init)). + apply fold_rec_aux. + simpl. red; intros; tauto. + simpl. red; intros. elim H0. + apply T.elements_keys_norepet. + apply H_base'. + simpl in H. red in H. apply H. red; intros. tauto. + Qed. + + End TREE_FOLD_IND. + + (** A nonnegative measure over trees *) + + Section MEASURE. + + Variable V: Type. + + Definition cardinal (x: T.t V) : nat := List.length (T.elements x). + + Theorem cardinal_remove: + forall x m y, T.get x m = Some y -> (cardinal (T.remove x m) < cardinal m)%nat. + Proof. + unfold cardinal; intros. + exploit T.elements_remove; eauto. intros (l1 & l2 & P & Q). + rewrite P, Q. rewrite ! app_length. simpl. omega. + Qed. + + Theorem cardinal_set: + forall x m y, T.get x m = None -> (cardinal m < cardinal (T.set x y m))%nat. + Proof. + intros. set (m' := T.set x y m). + replace (cardinal m) with (cardinal (T.remove x m')). + apply cardinal_remove with y. unfold m'; apply T.gss. + unfold cardinal. f_equal. apply T.elements_extensional. + intros. unfold m'. rewrite T.grspec, T.gsspec. + destruct (T.elt_eq i x); auto. congruence. + Qed. + + End MEASURE. + + (** Forall and exists *) + + Section FORALL_EXISTS. + + Variable A: Type. + + Definition for_all (m: T.t A) (f: T.elt -> A -> bool) : bool := + T.fold (fun b x a => b && f x a) m true. + + Lemma for_all_correct: + forall m f, + for_all m f = true <-> (forall x a, T.get x m = Some a -> f x a = true). + Proof. + intros m0 f. + unfold for_all. apply fold_rec; intros. + - (* Extensionality *) + rewrite H0. split; intros. rewrite <- H in H2; auto. rewrite H in H2; auto. + - (* Base case *) + split; intros. rewrite T.gempty in H0; congruence. auto. + - (* Inductive case *) + split; intros. + destruct (andb_prop _ _ H2). rewrite T.gsspec in H3. destruct (T.elt_eq x k). + inv H3. auto. + apply H1; auto. + apply andb_true_intro. split. + rewrite H1. intros. apply H2. rewrite T.gso; auto. congruence. + apply H2. apply T.gss. + Qed. + + Definition exists_ (m: T.t A) (f: T.elt -> A -> bool) : bool := + T.fold (fun b x a => b || f x a) m false. + + Lemma exists_correct: + forall m f, + exists_ m f = true <-> (exists x a, T.get x m = Some a /\ f x a = true). + Proof. + intros m0 f. + unfold exists_. apply fold_rec; intros. + - (* Extensionality *) + rewrite H0. split; intros (x0 & a0 & P & Q); exists x0; exists a0; split; auto; congruence. + - (* Base case *) + split; intros. congruence. destruct H as (x & a & P & Q). rewrite T.gempty in P; congruence. + - (* Inductive case *) + split; intros. + destruct (orb_true_elim _ _ H2). + rewrite H1 in e. destruct e as (x1 & a1 & P & Q). + exists x1; exists a1; split; auto. rewrite T.gso; auto. congruence. + exists k; exists v; split; auto. apply T.gss. + destruct H2 as (x1 & a1 & P & Q). apply orb_true_intro. + rewrite T.gsspec in P. destruct (T.elt_eq x1 k). + inv P. right; auto. + left. apply H1. exists x1; exists a1; auto. + Qed. + + Remark exists_for_all: + forall m f, + exists_ m f = negb (for_all m (fun x a => negb (f x a))). + Proof. + intros. unfold exists_, for_all. rewrite ! T.fold_spec. + change false with (negb true). generalize (T.elements m) true. + induction l; simpl; intros. + auto. + rewrite <- IHl. f_equal. + destruct b; destruct (f (fst a) (snd a)); reflexivity. + Qed. + + Remark for_all_exists: + forall m f, + for_all m f = negb (exists_ m (fun x a => negb (f x a))). + Proof. + intros. unfold exists_, for_all. rewrite ! T.fold_spec. + change true with (negb false). generalize (T.elements m) false. + induction l; simpl; intros. + auto. + rewrite <- IHl. f_equal. + destruct b; destruct (f (fst a) (snd a)); reflexivity. + Qed. + + Lemma for_all_false: + forall m f, + for_all m f = false <-> (exists x a, T.get x m = Some a /\ f x a = false). + Proof. + intros. rewrite for_all_exists. + rewrite negb_false_iff. rewrite exists_correct. + split; intros (x & a & P & Q); exists x; exists a; split; auto. + rewrite negb_true_iff in Q. auto. + rewrite Q; auto. + Qed. + + Lemma exists_false: + forall m f, + exists_ m f = false <-> (forall x a, T.get x m = Some a -> f x a = false). + Proof. + intros. rewrite exists_for_all. + rewrite negb_false_iff. rewrite for_all_correct. + split; intros. apply H in H0. rewrite negb_true_iff in H0. auto. rewrite H; auto. + Qed. + + End FORALL_EXISTS. + + (** More about [beq] *) + + Section BOOLEAN_EQUALITY. + + Variable A: Type. + Variable beqA: A -> A -> bool. + + Theorem beq_false: + forall m1 m2, + T.beq beqA m1 m2 = false <-> + exists x, match T.get x m1, T.get x m2 with + | None, None => False + | Some a1, Some a2 => beqA a1 a2 = false + | _, _ => True + end. + Proof. + intros; split; intros. + - (* beq = false -> existence *) + set (p1 := fun x a1 => match T.get x m2 with None => false | Some a2 => beqA a1 a2 end). + set (p2 := fun x a2 => match T.get x m1 with None => false | Some a1 => beqA a1 a2 end). + destruct (for_all m1 p1) eqn:F1; [destruct (for_all m2 p2) eqn:F2 | idtac]. + + cut (T.beq beqA m1 m2 = true). congruence. + rewrite for_all_correct in *. rewrite T.beq_correct; intros. + destruct (T.get x m1) as [a1|] eqn:X1. + generalize (F1 _ _ X1). unfold p1. destruct (T.get x m2); congruence. + destruct (T.get x m2) as [a2|] eqn:X2; auto. + generalize (F2 _ _ X2). unfold p2. rewrite X1. congruence. + + rewrite for_all_false in F2. destruct F2 as (x & a & P & Q). + exists x. rewrite P. unfold p2 in Q. destruct (T.get x m1); auto. + + rewrite for_all_false in F1. destruct F1 as (x & a & P & Q). + exists x. rewrite P. unfold p1 in Q. destruct (T.get x m2); auto. + - (* existence -> beq = false *) + destruct H as [x P]. + destruct (T.beq beqA m1 m2) eqn:E; auto. + rewrite T.beq_correct in E. + generalize (E x). destruct (T.get x m1); destruct (T.get x m2); tauto || congruence. + Qed. + + End BOOLEAN_EQUALITY. + + (** Extensional equality between trees *) + + Section EXTENSIONAL_EQUALITY. + + Variable A: Type. + Variable eqA: A -> A -> Prop. + Hypothesis eqAeq: Equivalence eqA. + + Definition Equal (m1 m2: T.t A) : Prop := + forall x, match T.get x m1, T.get x m2 with + | None, None => True + | Some a1, Some a2 => a1 === a2 + | _, _ => False + end. + + Lemma Equal_refl: forall m, Equal m m. + Proof. + intros; red; intros. destruct (T.get x m); auto. reflexivity. + Qed. + + Lemma Equal_sym: forall m1 m2, Equal m1 m2 -> Equal m2 m1. + Proof. + intros; red; intros. generalize (H x). destruct (T.get x m1); destruct (T.get x m2); auto. intros; symmetry; auto. + Qed. + + Lemma Equal_trans: forall m1 m2 m3, Equal m1 m2 -> Equal m2 m3 -> Equal m1 m3. + Proof. + intros; red; intros. generalize (H x) (H0 x). + destruct (T.get x m1); destruct (T.get x m2); try tauto; + destruct (T.get x m3); try tauto. + intros. transitivity a0; auto. + Qed. + + Instance Equal_Equivalence : Equivalence Equal := { + Equivalence_Reflexive := Equal_refl; + Equivalence_Symmetric := Equal_sym; + Equivalence_Transitive := Equal_trans + }. + + Hypothesis eqAdec: EqDec A eqA. + + Program Definition Equal_dec (m1 m2: T.t A) : { m1 === m2 } + { m1 =/= m2 } := + match T.beq (fun a1 a2 => proj_sumbool (a1 == a2)) m1 m2 with + | true => left _ + | false => right _ + end. + Next Obligation. + rename Heq_anonymous into B. + symmetry in B. rewrite T.beq_correct in B. + red; intros. generalize (B x). + destruct (T.get x m1); destruct (T.get x m2); auto. + intros. eapply proj_sumbool_true; eauto. + Qed. + Next Obligation. + assert (T.beq (fun a1 a2 => proj_sumbool (a1 == a2)) m1 m2 = true). + apply T.beq_correct; intros. + generalize (H x). + destruct (T.get x m1); destruct (T.get x m2); try tauto. + intros. apply proj_sumbool_is_true; auto. + unfold equiv, complement in H0. congruence. + Qed. + + Instance Equal_EqDec : EqDec (T.t A) Equal := Equal_dec. + + End EXTENSIONAL_EQUALITY. + + (** Creating a tree from a list of (key, value) pairs. *) + + Section OF_LIST. + + Variable A: Type. + + Let f := fun (m: T.t A) (k_v: T.elt * A) => T.set (fst k_v) (snd k_v) m. + + Definition of_list (l: list (T.elt * A)) : T.t A := + List.fold_left f l (T.empty _). + + Lemma in_of_list: + forall l k v, T.get k (of_list l) = Some v -> In (k, v) l. + Proof. + assert (REC: forall k v l m, + T.get k (fold_left f l m) = Some v -> In (k, v) l \/ T.get k m = Some v). + { induction l as [ | [k1 v1] l]; simpl; intros. + - tauto. + - apply IHl in H. unfold f in H. simpl in H. rewrite T.gsspec in H. + destruct H; auto. + destruct (T.elt_eq k k1). inv H. auto. auto. + } + intros. apply REC in H. rewrite T.gempty in H. intuition congruence. + Qed. + + Lemma of_list_dom: + forall l k, In k (map fst l) -> exists v, T.get k (of_list l) = Some v. + Proof. + assert (REC: forall k l m, + In k (map fst l) \/ (exists v, T.get k m = Some v) -> + exists v, T.get k (fold_left f l m) = Some v). + { induction l as [ | [k1 v1] l]; simpl; intros. + - tauto. + - apply IHl. unfold f; rewrite T.gsspec. simpl. destruct (T.elt_eq k k1). + right; econstructor; eauto. + intuition congruence. + } + intros. apply REC. auto. + Qed. + + Remark of_list_unchanged: + forall k l m, ~In k (map fst l) -> T.get k (List.fold_left f l m) = T.get k m. + Proof. + induction l as [ | [k1 v1] l]; simpl; intros. + - auto. + - rewrite IHl by tauto. unfold f; apply T.gso; intuition auto. + Qed. + + Lemma of_list_unique: + forall k v l1 l2, + ~In k (map fst l2) -> T.get k (of_list (l1 ++ (k, v) :: l2)) = Some v. + Proof. + intros. unfold of_list. rewrite fold_left_app. simpl. + rewrite of_list_unchanged by auto. unfold f; apply T.gss. + Qed. + + Lemma of_list_norepet: + forall l k v, list_norepet (map fst l) -> In (k, v) l -> T.get k (of_list l) = Some v. + Proof. + assert (REC: forall k v l m, + list_norepet (map fst l) -> + In (k, v) l -> + T.get k (fold_left f l m) = Some v). + { induction l as [ | [k1 v1] l]; simpl; intros. + contradiction. + inv H. destruct H0. + inv H. rewrite of_list_unchanged by auto. apply T.gss. + apply IHl; auto. + } + intros; apply REC; auto. + Qed. + + Lemma of_list_elements: + forall m k, T.get k (of_list (T.elements m)) = T.get k m. + Proof. + intros. destruct (T.get k m) as [v|] eqn:M. + - apply of_list_norepet. apply T.elements_keys_norepet. apply T.elements_correct; auto. + - destruct (T.get k (of_list (T.elements m))) as [v|] eqn:M'; auto. + apply in_of_list in M'. apply T.elements_complete in M'. congruence. + Qed. + + End OF_LIST. + + Lemma of_list_related: + forall (A B: Type) (R: A -> B -> Prop) k l1 l2, + list_forall2 (fun ka kb => fst ka = fst kb /\ R (snd ka) (snd kb)) l1 l2 -> + option_rel R (T.get k (of_list l1)) (T.get k (of_list l2)). + Proof. + intros until k. unfold of_list. + set (R' := fun ka kb => fst ka = fst kb /\ R (snd ka) (snd kb)). + set (fa := fun (m : T.t A) (k_v : T.elt * A) => T.set (fst k_v) (snd k_v) m). + set (fb := fun (m : T.t B) (k_v : T.elt * B) => T.set (fst k_v) (snd k_v) m). + assert (REC: forall l1 l2, list_forall2 R' l1 l2 -> + forall m1 m2, option_rel R (T.get k m1) (T.get k m2) -> + option_rel R (T.get k (fold_left fa l1 m1)) (T.get k (fold_left fb l2 m2))). + { induction 1; intros; simpl. + - auto. + - apply IHlist_forall2. unfold fa, fb. rewrite ! T.gsspec. + destruct H as [E F]. rewrite E. destruct (T.elt_eq k (fst b1)). + constructor; auto. + auto. } + intros. apply REC; auto. rewrite ! T.gempty. constructor. + Qed. + + + (* folding a commutative function respects permutations. *) + + Require Import Permutation. + + (* norepet is from Compcert's Coqlib, NoDup is from the Coq standard lib, they mean the same thing. *) + Lemma norepet_NoDup : forall A (l : list A), list_norepet l -> NoDup l. + Proof. + induction 1; constructor; auto. + Qed. + + Lemma NoDup_map : forall A B (f:A->B) (l : list A), NoDup (map f l) -> NoDup l. + Proof. + intros. + remember (map f l) as l1. + revert l Heql1. + induction H. + + destruct l; simpl in *; try congruence; constructor. + + intros. + destruct l0; simpl in *; try congruence; constructor; + inversion Heql1; subst; auto using in_map. + Qed. + + Lemma elements_NoDup : forall A (m : T.t A), NoDup (T.elements m). + Proof. + intros. + apply NoDup_map with (f:=fst). + apply norepet_NoDup. + apply T.elements_keys_norepet. + Qed. + + Lemma set_permutation : forall {A} k (v:A) m, + T.get k m = None -> + Permutation (T.elements (T.set k v m)) ((k,v) :: T.elements m). + Proof. + intros. + apply NoDup_Permutation. + - apply elements_NoDup. + - constructor. + intros HIn; apply T.elements_complete in HIn; congruence. + apply elements_NoDup. + - intros [k' v']. split. + * destruct (T.elt_eq k' k). + + subst. + intros HIn. + apply T.elements_complete in HIn. + rewrite T.gss in HIn. + inversion HIn; subst. + simpl; auto. + + intros HIn. + apply T.elements_complete in HIn. + rewrite T.gso in HIn by auto. + right. + apply T.elements_correct. + auto. + * destruct 1. + + inversion H0; subst. + apply T.elements_correct. + rewrite T.gss. auto. + + assert (k'<>k). + { apply T.elements_complete in H0. intros ?. congruence. } + apply T.elements_correct. + apply T.elements_complete in H0. + rewrite T.gso; auto. + Qed. + + Lemma get_permutation : forall {A} k (v:A) m, + T.get k m = Some v -> + exists lst, Permutation (T.elements m) (lst ++ (k,v)::nil). + Proof. + intros A k v m Hget. + assert (Hin := T.elements_correct Hget). + assert (Hnodup := T.elements_keys_norepet m). + apply norepet_NoDup in Hnodup. + destruct (in_split _ _ Hin) as [elements1 [ elements2 elements_eq]]. + exists (elements1 ++ elements2). + rewrite elements_eq in *. + clear Hin elements_eq. + rewrite <- !app_assoc. + apply Permutation_app_head. + apply Permutation_cons_append. + Qed. + + Lemma NoDup_snoc : forall {A} (x:A) l, + NoDup (x::l) -> NoDup (l++x::nil). + Proof. + induction l. + - constructor. + simpl; intuition. + constructor. + - simpl. + inversion 1. + inversion H3. + simpl in *. + subst. + constructor. + + rewrite in_app. + simpl in *. + intuition. + + apply IHl. + constructor; + intuition. + Qed. + + + Section fold_permutation. + Context (A B K : Type) + (f1 : A->B->A) + (f1_comm : forall (b1 b2 : B) a, f1 (f1 a b1) b2 = f1 (f1 a b2) b1) + (f : A -> K -> B -> A) + (f_comm : forall k1 b1 k2 b2 a, + k1 <> k2 -> + f (f a k1 b1) k2 b2 = f (f a k2 b2) k1 b1). + + Lemma fold1_permutation : forall l l', + Permutation l l' -> forall init, + List.fold_left f1 l init = List.fold_left f1 l' init. + Proof. + induction 1; intros. + - reflexivity. + - simpl. rewrite IHPermutation. + reflexivity. + - simpl. rewrite f1_comm. reflexivity. + - rewrite IHPermutation1, IHPermutation2. + reflexivity. + Qed. + + Definition NoDup_keys (l : list (K*B)) := NoDup (map fst l). + + Lemma Permutation_NoDup_keys : forall l l', + Permutation l l' -> NoDup_keys l -> NoDup_keys l'. + Proof. + unfold NoDup_keys. + intros. + apply Permutation_NoDup with (map fst l). + - apply Permutation_map. + assumption. + - assumption. + Qed. + + Lemma NoDup_keys_inv_cons : forall a l, + NoDup_keys (a::l) -> + NoDup_keys l. + Proof. + unfold NoDup_keys. + simpl. + intros. + rewrite NoDup_cons_iff in H. + tauto. + Qed. + + Lemma fold_permutation : forall l l', + Permutation l l' -> NoDup_keys l -> forall init, + List.fold_left (fun a p => f a (fst p) (snd p)) l init + = List.fold_left (fun a p => f a (fst p) (snd p)) l' init. + Proof. + induction 1; intros. + - reflexivity. + - simpl. rewrite IHPermutation. + + reflexivity. + + apply NoDup_keys_inv_cons in H0. + assumption. + - simpl. + rewrite f_comm. + + reflexivity. + + unfold NoDup_keys in H. + simpl in H. + rewrite !NoDup_cons_iff in H. + simpl in H. + intros H1. + symmetry in H1. + tauto. + - rewrite IHPermutation1, IHPermutation2 by eauto using Permutation_NoDup_keys . + reflexivity. + Qed. + End fold_permutation. + + Section fold1_set. + Context {A} {B} (f : A->B->A) + (f_comm : forall (x y : B) a, f (f a x) y = f (f a y) x). + + Lemma f1_comm : forall (x y : T.elt * B) (a : A), + (fun a0 p => f a0 (snd p)) + ((fun a0 p => f a0 (snd p)) a x) y = + (fun a0 p => f a0 (snd p)) + ((fun a0 p => f a0 (snd p)) a y) x. + intros [kx x] [ky y] init. + simpl. + apply f_comm. + Qed. + + Lemma fold1_extensional : forall init m m', + (forall k, T.get k m = T.get k m') -> + T.fold1 f m init = T.fold1 f m' init. + Proof. + intros. + repeat rewrite T.fold1_spec. + refine (fold1_permutation _ _ _ _). + + intros. + apply f_comm. + + apply NoDup_Permutation; try apply elements_NoDup. + intros [k' v']. + split; intros HIn; + apply T.elements_complete in HIn; + apply T.elements_correct; + rewrite H in *; + auto. + Qed. + + Lemma elements_set_decompose : forall {A} k (v:A) m, + T.elements (T.set k v m) = T.elements (T.set k v (T.remove k m)). + Proof. + intros. + apply T.elements_extensional. + intros i. destruct (T.elt_eq i k). + + subst. repeat rewrite T.gss. reflexivity. + + repeat rewrite T.gso by auto. rewrite T.gro by auto. + reflexivity. + Qed. + + Lemma fold1_set : forall init k v m, + T.fold1 f (T.set k v m) init = T.fold1 f (T.remove k m) (f init v). + Proof. + intros. + repeat rewrite T.fold1_spec. + rewrite elements_set_decompose. + assert (HPerm: Permutation (T.elements (T.set k v (T.remove k m))) + ((k,v):: T.elements (T.remove k m))) + by (apply set_permutation; apply T.grs). + rewrite (fold1_permutation _ f1_comm HPerm). + reflexivity. + Qed. + + Lemma fold1_get : forall init k v m, + T.get k m = Some v -> + T.fold1 f m init = T.fold1 f (T.remove k m) (f init v). + Proof. + intros. + replace (T.fold1 f m init) + with (T.fold1 f (T.set k v m) init). + + apply fold1_set. + + apply fold1_extensional. + intros k0. + destruct (T.elt_eq k0 k); subst. + - rewrite T.gss; auto. + - rewrite T.gso; auto. + Qed. + + + Lemma fold1_remove_set : forall init k v m, + T.fold1 f (T.remove k (T.set k v m)) init = T.fold1 f (T.remove k m) init. + Proof. + intros. + apply fold1_extensional. + intros k0. + repeat rewrite T.grspec. + destruct (T.elt_eq k0 k); subst. + + reflexivity. + + rewrite T.gso; auto. + Qed. + End fold1_set. + + Section sum. + Require Import ZArith. + Definition sum (m : T.t Z) := T.fold1 Z.add m 0. + + Lemma plus_comm : forall x y a : Z, a + x + y = a + y + x. + intros; omega. + Qed. + + Theorem constant_sum : forall (m : T.t Z) k v k' v' i, + T.get k m = Some v -> + T.get k' m = Some v' -> + k <> k' -> + sum (T.set k (v - i) (T.set k' (v' + i) m)) = sum m. + Proof. + intros m k v k' v' i Hget_k Hget_k' Hneq. + unfold sum. + rewrite fold1_get with (k0:=k') (v0:=(v'+i)); + [|apply plus_comm|rewrite T.gso by congruence; rewrite T.gss; reflexivity]. + rewrite fold1_get with (k0:=k) (v0:=(v - i)); + [|apply plus_comm|rewrite T.gro by congruence; rewrite T.gss; reflexivity]. + rewrite fold1_get with (k0:=k') (v0:=v') (m0:=m); + [|apply plus_comm|congruence]. + rewrite fold1_get with (k0:=k) (v0:=v) (m0:=T.remove k' m); + [|apply plus_comm|rewrite T.gro; congruence]. + replace (T.fold1 Z.add + (T.remove k (T.remove k' (T.set k (v - i) (T.set k' (v' + i) m)))) + (0 + (v' + i) + (v - i))) + with (T.fold1 Z.add + (T.remove k (T.remove k' m)) + (0 + (v' + i) + (v - i))). + + f_equal. + omega. + + apply fold1_extensional. + - apply plus_comm. + - intros k0. + destruct (T.elt_eq k0 k); destruct (T.elt_eq k0 k'); subst; + (repeat first [rewrite T.grs| rewrite T.gro by auto|rewrite T.gso by auto ]); + reflexivity. + Qed. + + Definition nonzero_elements (m : T.t Z) := + List.filter (fun e => negb (Z.eqb (snd e) 0)) (T.elements m). + + Lemma nonzero_elements_spec : forall m k v, + In (k,v) (nonzero_elements m) <-> + (In (k,v) (T.elements m) /\ v<>0). + Proof. + unfold nonzero_elements. + intros. + rewrite filter_In. + rewrite negb_true_iff. + simpl. + rewrite Z.eqb_neq. + intuition. + Qed. + + Lemma sum_fold_left_nonzero : forall m init, + List.fold_left (fun a p => a + (snd p)) (nonzero_elements m) init = + List.fold_left (fun a p => a + (snd p)) (T.elements m) init. + Proof. + intros m. + unfold nonzero_elements. + induction (T.elements m). + - reflexivity. + - simpl. + destruct (snd a =? 0) eqn:?. + + rewrite Z.eqb_eq in Heqb. + simpl. + intros init. + replace (init + snd a) with init by omega. + apply IHl. + + intros init. + simpl. + apply IHl. + Qed. + + Lemma NoDup_filter : forall A f (l:list A), NoDup l -> NoDup (filter f l). + Proof. + induction l. + + simpl; constructor. + + intros H. + rewrite NoDup_cons_iff in H. + destruct H. + simpl. + destruct (f a). + * apply NoDup_cons. + rewrite filter_In; tauto. + auto. + * auto. + Qed. + + Lemma nonzero_elements_extensional: + forall (m n: T.t Z), + (forall i, T.get_default 0 i m = T.get_default 0 i n) -> + Permutation (nonzero_elements m) (nonzero_elements n). + Proof. + intros. + apply NoDup_Permutation. + - apply NoDup_filter. + apply elements_NoDup. + - apply NoDup_filter. + apply elements_NoDup. + - intros [k v]. + rewrite ?nonzero_elements_spec. + unfold T.get_default in H. + intuition. + + apply T.elements_complete in H1. + apply T.elements_correct. + specialize (H k). + rewrite H1 in H. + destruct (T.get k n) as [v'|]; [congruence|subst;tauto]. + + apply T.elements_complete in H1. + apply T.elements_correct. + specialize (H k). + rewrite H1 in H. + destruct (T.get k m) as [v'|]; [congruence|subst;tauto]. + Qed. + + Lemma sum_extensional: + forall (m n: T.t Z) init, + (forall i, T.get_default 0 i m = T.get_default 0 i n) -> + T.fold1 Z.add m init = T.fold1 Z.add n init. + Proof. + intros. + rewrite ?T.fold1_spec. + rewrite <- ?sum_fold_left_nonzero. + apply fold1_permutation. + - intros; omega. + - apply nonzero_elements_extensional. + auto. + Qed. + + Lemma get_default_ss : forall A (def:A) k v m, + T.get_default def k (T.set k v m) = v. + Proof. + intros. + unfold T.get_default. + rewrite T.gss. + reflexivity. + Qed. + + Lemma get_default_so : forall A (def:A) k k0 v m, + k <> k0 -> + T.get_default def k (T.set k0 v m) = T.get_default def k m. + Proof. + intros. + unfold T.get_default. + rewrite T.gso; auto. + Qed. + + Lemma get_default_ro : forall A (def:A) k k0 m, + k <> k0 -> + T.get_default def k (T.remove k0 m) = T.get_default def k m. + Proof. + intros. + unfold T.get_default. + rewrite T.gro; auto. + Qed. + + Lemma get_default_rs : forall A (def:A) k m, + T.get_default def k (T.remove k m) = def. + Proof. + intros. + unfold T.get_default. + rewrite T.grs; auto. + Qed. + + Lemma sum_get_default : forall init k v m, + T.get_default 0 k m = v -> + T.fold1 Z.add m init = T.fold1 Z.add (T.remove k m) (Z.add init v). + Proof. + intros. + replace (T.fold1 Z.add m init) + with (T.fold1 Z.add (T.set k v m) init). + + apply fold1_set. + intros; omega. + + apply sum_extensional. + intros k0. + destruct (T.elt_eq k0 k); subst. + - rewrite get_default_ss; auto. + - rewrite get_default_so; auto. + Qed. + + (* A fancier version of the theorem constant_sum, which knows about + reading the default value zero. *) + Theorem constant_sum' : forall (m : T.t Z) k v k' v' i, + T.get_default 0 k m = v -> + T.get_default 0 k' m = v' -> + k <> k' -> + sum (T.set k (v - i) (T.set k' (v' + i) m)) = sum m. + Proof. + intros m k v k' v' i Hget_k Hget_k' Hneq. + unfold sum. + rewrite fold1_get with (k0:=k') (v0:=(v'+i)); + [|apply plus_comm|rewrite T.gso by congruence; rewrite T.gss; reflexivity]. + rewrite fold1_get with (k0:=k) (v0:=(v - i)); + [|apply plus_comm|rewrite T.gro by congruence; rewrite T.gss; reflexivity]. + + rewrite sum_get_default with (k:=k') (v:=v') (m:=m); [|assumption]. + rewrite sum_get_default with (k:=k) (v:=v) (m:=T.remove k' m); [|rewrite get_default_ro; auto]. + replace (T.fold1 Z.add + (T.remove k (T.remove k' (T.set k (v - i) (T.set k' (v' + i) m)))) + (0 + (v' + i) + (v - i))) + with (T.fold1 Z.add + (T.remove k (T.remove k' m)) + (0 + (v' + i) + (v - i))). + + f_equal. + omega. + + apply fold1_extensional. + - intros; omega. + - intros k0. + destruct (T.elt_eq k0 k); destruct (T.elt_eq k0 k'); subst; + (repeat first [rewrite T.grs| rewrite T.gro by auto|rewrite T.gso by auto ]); + reflexivity. + Qed. + + Lemma sum_nonnegative' : forall (m : T.t Z) init, + (forall k v, T.get k m = Some v -> v >= 0) -> + T.fold1 Z.add m init >= init. + Proof. + intros. + rewrite T.fold1_spec. + assert (H' : forall k v, In (k,v) (T.elements m) -> v >= 0). + { + intros. + apply T.elements_complete in H0. + eauto. + } + clear H. + revert init H'. + induction (T.elements m); intros; simpl in *. + - omega. + - apply Zge_trans with (init + snd a). + apply IHl. + + intros; apply H' with k. auto. + + assert (snd a >= 0). + { destruct a as [k v]. apply H' with k. left. reflexivity. } + omega. + Qed. + + Lemma sum_nonnegative : forall (m : T.t Z), + (forall k v, T.get k m = Some v -> v >= 0) -> + sum m >= 0. + Proof. + intros. + unfold sum. + apply sum_nonnegative'; auto; omega. + Qed. + + Lemma sum_bound1 : forall (m : T.t Z) k1 B, + (forall k v, T.get k m = Some v -> v >= 0) -> + sum m <= B -> + T.get_default 0 k1 m <= B. + Proof. + intros. + assert (B >= 0). + { + pose (sum_nonnegative H). + omega. + } + + unfold sum in H0. + rewrite (@sum_get_default _ k1 _ _ eq_refl) in H0. + assert (forall (k : T.elt) (v : Z), + T.get k (T.remove k1 m) = Some v -> v >= 0). + { + intros. + destruct (T.elt_eq k k1); subst. + - rewrite T.grs in H2; congruence. + - rewrite !T.gro in H2 by congruence. + eapply H; eauto. + } + assert (H4 := @sum_nonnegative' (T.remove k1 m) + (0 + T.get_default 0 k1 m) + H2). + omega. + Qed. + + Lemma sum_bound2 : forall (m : T.t Z) k1 k2 B, + k1<>k2 -> + (forall k v, T.get k m = Some v -> v >= 0) -> + sum m <= B -> + T.get_default 0 k1 m + T.get_default 0 k2 m <= B. + Proof. + intros. + assert (B >= 0). + { + pose (sum_nonnegative H0). + omega. + } + + unfold sum in H1. + rewrite (@sum_get_default _ k1 _ _ eq_refl) in H1. + rewrite (@sum_get_default _ k2 _ _ eq_refl) in H1. + assert (forall (k : T.elt) (v : Z), + T.get k (T.remove k2 (T.remove k1 m)) = Some v -> v >= 0). + { + intros. + destruct (T.elt_eq k k2), (T.elt_eq k k1); subst. + - rewrite T.grs in H3; congruence. + - rewrite T.grs in H3. congruence. + - rewrite !T.gro, T.grs in H3 by congruence. congruence. + - rewrite !T.gro in H3 by congruence. + eapply H0; eauto. + } + assert (H4 := @sum_nonnegative' (T.remove k2 (T.remove k1 m)) + (0 + T.get_default 0 k1 m + T.get_default 0 k2 (T.remove k1 m)) + H3). + remember ( T.fold1 Z.add (T.remove k2 (T.remove k1 m)) + (0 + T.get_default 0 k1 m + T.get_default 0 k2 (T.remove k1 m))) as blah. + clear Heqblah. + rewrite get_default_ro in H4 by congruence. + omega. + Qed. + + Lemma sum_swap : forall (m : T.t Z) k v k' v', + k <> k' -> + sum (T.set k v (T.set k' v' m)) = + sum (T.set k' v' (T.set k v m)). + Proof. + intros. + unfold sum. + apply fold1_extensional. + - intros; omega. + - intros k0. + destruct (T.elt_eq k0 k); destruct (T.elt_eq k0 k'); subst; + try congruence; + (repeat first [rewrite T.gss| rewrite T.grs| rewrite T.gro by auto|rewrite T.gso by auto ]; + auto). + Qed. + + End sum. +End Tree_Properties. + +Module PTree_Properties := Tree_Properties(PTree). +Module ZTree_Properties := Tree_Properties(ZTree). +Module Int256Tree_Properties := Tree_Properties(Int256Tree). + +(* Ideally this would be a module functor, but I'm too lazy to figure out + how to make it depend on both TREE and MAP at the same time, + so I just prove it for Int256Map & Int256Tree for now. *) +Section MapOfTree. + Context (A B : Type) (f : A -> B) (default : A). + + + Definition map_of_tree (t : Int256Tree.t A) : Int256Map.t B := + Int256Tree.fold (fun m k v => Int256Map.set k (f v) m) + t + (Int256Map.init (f default)). + + Lemma mot_ins_commute : + let f := (fun m k v => Int256Map.set k (f v) m) + in + forall k1 b1 k2 b2 a, + k1 <> k2 -> + f (f a k1 b1) k2 b2 = f (f a k2 b2) k1 b1. + Proof. + intros. + unfold f0. + unfold Int256Map.set. + unfold PMap.set. + simpl. + rewrite PTree.set_swap. + - reflexivity. + - intros C; apply Int256Indexed.index_inj in C. + congruence. + Qed. + + Lemma fold_left_last : forall A B (f:A->B->A) l x init, + fold_left f (l ++ x :: nil) init = f (fold_left f l init) x. + Proof. + intros. + rewrite fold_left_app. + reflexivity. + Qed. + + Require Import Permutation. + + Lemma map_of_tree_Some : forall i t v, + Int256Tree.get i t = Some v -> + Int256Map.get i (map_of_tree t) = f v. + Proof. + intros. + unfold map_of_tree. + rewrite Int256Tree.fold_spec. + assert (Hin := Int256Tree.elements_correct _ _ H). + assert (Hnodup := Int256Tree.elements_keys_norepet t). + apply Int256Tree_Properties.norepet_NoDup in Hnodup. + destruct (in_split _ _ Hin) as [elements1 [ elements2 elements_eq]]. + rewrite elements_eq in *. + clear Hin elements_eq. + rewrite (@Int256Tree_Properties.fold_permutation _ _ _ + (fun m k v => Int256Map.set k (f v) m) + mot_ins_commute + (elements1 ++ (i, v) :: elements2) + (elements1 ++ elements2 ++ (i, v) :: nil)). + - rewrite !app_assoc. + rewrite fold_left_last. + simpl. + rewrite Int256Map.gss. + reflexivity. + - apply Permutation_app_head. + apply Permutation_cons_append. + - exact Hnodup. + Qed. + + Lemma map_of_tree_None' : forall i lst m, + ~In i (map fst lst) -> + Int256Map.get i m = f default -> + Int256Map.get i + (fold_left + (fun a p => + Int256Map.set (fst p) (f (snd p)) a) lst + m) + = f default. + Proof. + induction lst. + - simpl. + auto. + - simpl. + intros. + apply IHlst. + tauto. + rewrite Int256Map.gso; auto; tauto. + Qed. + + Lemma map_of_tree_None : forall i t, + Int256Tree.get i t = None -> + Int256Map.get i (map_of_tree t) = f default. + Proof. + intros. + unfold map_of_tree. + rewrite Int256Tree.fold_spec. + rewrite map_of_tree_None'. + - auto. + - intros Hin. + apply list_in_map_inv in Hin. + destruct Hin as [[v ty] [v_eq Hin]]. + apply Int256Tree.elements_complete in Hin. + simpl in v_eq; inversion v_eq; subst. + congruence. + - rewrite Int256Map.gi; auto. + Qed. + + Lemma map_of_tree_set : forall i t v, + map_of_tree (Int256Tree.set i v t) + = Int256Map.set i (f v) (map_of_tree t). + Proof. + (* I'm pretty sure this is true, but proving it seems annoying. *) + Admitted. + +End MapOfTree. + + + +(** * Useful notations *) + +Notation "a ! b" := (PTree.get b a) (at level 1). +Notation "a !! b" := (PMap.get b a) (at level 1). diff --git a/src/core/Cval.v b/src/core/Cval.v new file mode 100755 index 0000000..901be08 --- /dev/null +++ b/src/core/Cval.v @@ -0,0 +1,429 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** Defining the "extended C value" [cval], its operations, and its relation + with the vanilla C value type [val]. *) +(** This file defines a cval layer above CompCert C memory model. It also + * defines some operators for cval, and proves simulation relation between + * these two layers for each operator. *) + +(* Standard library modules *) +Require Import BinInt. (* Z_scope, (_ | _) *) + +Require Import cclib.Coqlib. +Require Import cclib.Integers. +Require Import cclib.Maps. +Require Import cclib.Errors. +Require Import backend.Cop. +Require Import backend.Ctypes. +Require Import backend.Values.HighValues. +Require Import backend.Environments.ExtEnv. + +Notation tint := (Tint I256 Unsigned). +Notation tchar := (Tint I8 Unsigned). +Notation tchar_pointer := (Tpointer tchar). + +Inductive cval : Type := + | CVval : val -> cval + | CVany : cval (* corresponds to any C value, can be refined by any, even undefined value *) + | CVstruct : cstruct_val -> cval + | CVarray : carray_val -> cval + | CVhashmap : chashmap_val -> cval + +with cstruct_val : (*fieldlist ->*) Type := + | CSmap (map : PTree.t cval) + +with carray_val : Type := + | CAmap (map : ZMap.t cval) + +with chashmap_val : Type := + | CHmap (map : Int256Map.t cval) +. + +Definition cval_unary_operation op (v: cval) t : option cval := + match v with + | CVval v' => match sem_unary_operation op v' t with + | Some v'' => Some (CVval v'') + | None => None + end + | _ => None + end. + +Definition cval_binary_operation (*`{Mem.MemoryModelOps}*) + op cv1 t1 cv2 t2 : option cval := + match cv1, cv2 with + | CVval v1, CVval v2 => + (*if is_pointer_comparison op t1 t2 then None + else*) + match sem_binary_operation op v1 t1 v2 t2 (*Mem.empty*) with + | Some v => Some (CVval v) + | None => None + end + | _, _ => None + end. + +Definition cval_field_access f cv t : option cval := + match cv, t with + | CVstruct (CSmap map), Tstruct _ flds => + match struct_field flds f with + | Some _ => PTree.get f map + | _ => None + end + | _, _ => None + end. + +Definition cval_field_update f cv t newfield : option cval := + match cv, t with + | CVstruct (CSmap map), Tstruct _ flds => + (* match field_offset f flds with + | OK _ => *) Some (CVstruct (CSmap (PTree.set f newfield map))) +(* | _ => None + end*) + | _, _ => None + end. + +Definition cval_array_indexing idx cv t : option cval := + match cv, t with + | CVarray (CAmap map), Tarray _ n => + if zle 0 idx && zlt idx n then + Some (ZMap.get idx map) + else + None + | _, _ => None + end. + +Definition cval_array_update idx cv t newelem : option cval := + match cv, t with + | CVarray (CAmap map), Tarray _ n => + if zle 0 idx && zlt idx n then + Some (CVarray (CAmap (ZMap.set idx newelem map))) + else + None + | _, _ => None + end. + +Definition cval_hashmap_lookup idx cv t : option cval := + match cv, t with + | CVhashmap (CHmap map), Thashmap tint _ => + Some (Int256Map.get idx map) + | _, _ => None + end. + +Definition cval_hashmap_update idx cv t newelem : option cval := + match cv, t with + | CVhashmap (CHmap map), Thashmap tint n => + Some (CVhashmap (CHmap (Int256Map.set idx newelem map))) + | _, _ => None + end. + +Inductive cval_basic : cval -> Prop := + | CVBint : forall i, cval_basic (CVval (Vint i)) + . + +Inductive cval_basic_match : val -> cval -> Prop := +| CVBMval_int : forall i, cval_basic_match (Vint i) (CVval (Vint i)) +| CVBMany : forall v, cval_basic_match v CVany + . + + Inductive cval_match : val -> cval -> Prop := + | CVMval : forall v, cval_match v (CVval v) + (* ^^ In the C backend there is an extra condition that if v is a pointer + the identifier must be less than "glob_threshold", 100000. *) + | CVMany : forall v, cval_match v CVany. + + Definition is_immediate_type (ty : type) : bool := + match ty with + | Tint _ _ => true + | _ => false + end. + + Inductive cval_match_indirect (m : ext_env) + : type -> ident_ext -> cval -> Prop := + | CVMIval : forall ty b v cv, + is_immediate_type ty = true -> + IdentExtMap.get b m = EEVal v -> + cval_match v cv -> + cval_match_indirect m ty b cv + | CVMIstruct : forall i flds b vs, + cval_match_struct m flds b vs -> + cval_match_indirect m (Tstruct i flds) b (CVstruct vs) + | CVMIarray : forall ty n b vs, + cval_match_array m ty b vs n -> + cval_match_indirect m (Tarray ty n) b (CVarray vs) + | CVMIhash : forall ty b vs, + cval_match_hash_int m ty b vs -> + cval_match_indirect m (Thashmap tint ty) b (CVhashmap vs) + + with cval_match_struct (m : ext_env) + : fieldlist -> ident_ext -> cstruct_val -> Prop := + | CFMmapped : forall flds b map, (* struct's head address is b *) + (forall i ofs ty, + struct_field flds i = Some (ofs, ty) -> (* ofs is just the count number of field. TODO: maybe we don't need it at all? *) + exists cv, + PTree.get i map = Some cv /\ + cval_match_indirect m ty (Field b i) cv) -> + cval_match_struct m flds b (CSmap map) + + with cval_match_array (m : ext_env) + : type -> ident_ext -> carray_val -> Z -> Prop := + | CAMmapped : forall ty b map n, + (forall idx, 0 <= idx < n -> + cval_match_indirect m ty (Index b (Int256.repr idx)) + (ZMap.get idx map)) -> + cval_match_array m ty b (CAmap map) n + + with cval_match_hash_int (m : ext_env) + : type -> ident_ext -> chashmap_val -> Prop := + | CHIMmapped : forall ty b map, + (forall idx, + cval_match_indirect m ty (Index b idx) + (Int256Map.get idx map)) -> + cval_match_hash_int m ty b (CHmap map) . + + Lemma cval_basic_match_cval_match v cv : + cval_basic_match v cv -> cval_match v cv. + Proof. + inversion 1; constructor; reflexivity. + Qed. + + (* v1 ---sem_unary---> v2 + | | + cval_match cval_match + | | + cv1 --cval_unary--> cv2 *) + Lemma cval_sem_unary_operation op v1 cv1 cv2 t : + cval_match v1 cv1 -> + cval_unary_operation op cv1 t = Some cv2 -> + exists v2, sem_unary_operation op v1 t = Some v2 /\ cval_match v2 cv2. + Proof. + inversion 1 as [v1i | ?]; subst; clear H. + - destruct (sem_unary_operation op v1 t) as [ vi |] eqn:unary_eq; try discriminate. + + simpl. + rewrite unary_eq. + injection 1 as <-. + eexists. + split; [reflexivity|]. + constructor. + + simpl. + rewrite unary_eq. + intros; congruence. + - destruct (sem_unary_operation op v1 t) as [ vi |] eqn:unary_eq; try discriminate. + Qed. + + Ltac TrivialExists := + match goal with + | |- exists v', Some ?v = Some v' /\ _ => exists v; split; auto + | |- exists v', (forall _, Some ?v = Some v') /\ _ => exists v; split; auto + | _ => idtac + end. + + (* v1,v2 ---sem_binary---> v + | | + cval_match cval_match + | | + cv1,cv2 --cval_binary--> cv *) + Lemma cval_sem_binary_operation op v1 v2 cv1 cv2 cv t1 t2 : + cval_match v1 cv1 -> cval_match v2 cv2 -> + cval_binary_operation op cv1 t1 cv2 t2 = Some cv -> + exists v, sem_binary_operation op v1 t1 v2 t2 = Some v /\ + cval_match v cv. + Proof. + inversion 1 as [ v1i | ]; + inversion 1 as [ v2i | ]; + subst; try discriminate; clear H H2. + unfold cval_binary_operation. + destruct cv as [ vi | | | |]; + destruct (sem_binary_operation op v1 t1 v2 t2) eqn:binary_eq; + try discriminate; + injection 1 as ->. + TrivialExists. + constructor. + Qed. + + (* cv1 ---cval_field_access f---> cv2 + | | + match_indirect match_indirect + | | + t1, b, ofs --access field f--> t2, b, (ofs + delta) *) + Lemma cval_sem_field_access f cv1 cv2 t1 : + cval_field_access f cv1 t1 = Some cv2 -> + exists id flds t2 delta, + t1 = Tstruct id flds /\ + struct_field flds f = Some (delta, t2) /\ + forall m b, + cval_match_indirect m t1 b cv1 -> + cval_match_indirect m t2 (Field b f) cv2. + Proof. + unfold cval_field_access. + destruct cv1 as [| | | [] |]; try discriminate. + destruct c as [map]. + destruct t1 as [| | | | | | tid flds | |]; simpl; try discriminate. + destruct (struct_field flds f) as [[delta t2]|] eqn:offset_eq; try discriminate. + intros cv2_eq. + + exists tid, flds, t2, delta. + repeat split; try assumption. + + intros m b match_indir. + inversion match_indir as [| ? ? ? ? match_struct | |]; + try match goal with [ H: is_immediate_type _ = true |- _] => now inversion H end; subst. + inversion match_struct as [ ? ? ? match_field ]; subst. + destruct (match_field _ _ _ offset_eq) + as (cv' & cv'_eq & match_indir'). + rewrite cv2_eq in cv'_eq. + injection cv'_eq as ->. + exact match_indir'. + Qed. + + Require Import lib.ArithInv. + + (* cv1 ---cval_array_indexing idx---> cv2 + | | + match_indirect match_indirect + | | + t1, b, ofs -- access index idx --> t2, b, (ofs + idx * sizeof t2) *) + Lemma cval_sem_array_indexing idx cv1 cv2 t1 : + cval_array_indexing idx cv1 t1 = Some cv2 -> + exists t2 n, + t1 = Tarray t2 n /\ + forall m b, + cval_match_indirect m t1 b cv1 -> + cval_match_indirect m t2 (Index b (Int256.repr idx)) cv2. + Proof. + unfold cval_array_indexing. + destruct cv1 as [| | | | []]; try discriminate. + destruct c as [map]. + destruct t1 as [| | | | | t2 n | | |]; try discriminate. + destruct (zle 0 idx && zlt idx z) eqn:bound_check; try discriminate. + intros cv2_eq; injection cv2_eq as <-. + exists t1, z. + repeat split; try assumption. + + intros m b match_indir. + inversion match_indir as [| | ? ? ? ? match_array |]; + try match goal with [ H: is_immediate_type _ = true |- _] => now inversion H end; subst. + inversion match_array as [ ? ? ? ? ? bounded match_index ]; subst. + apply H. + inv_arith. + split; auto. + Qed. + + Lemma cval_sem_hashmap_lookup idx cv1 cv2 t1 : + cval_hashmap_lookup idx cv1 t1 = Some cv2 -> + exists t2, + t1 = Thashmap tint t2 /\ + forall m b, + cval_match_indirect m t1 b cv1 -> + cval_match_indirect m t2 (Index b idx) cv2. + Proof. + unfold cval_hashmap_lookup. + destruct cv1 as [| | | | []]; try discriminate. + (* destruct c as [map]. *) + destruct t1 as [| | | | t1 t2 | | | | ]; try discriminate. + destruct t1; try discriminate. + destruct i; try discriminate. + destruct s; try discriminate. + intros cv2_eq; injection cv2_eq as <-. + exists t2. + repeat split; try assumption. + + intros m b match_indir. + inversion match_indir; + try match goal with [ H: is_immediate_type _ = true |- _] => now inversion H end; subst. + inversion H2 as [ ? ? ? match_hash]; subst. + apply match_hash. + Qed. + + Lemma cval_match_indirect_eq m m' i ty cv : + cval_match_indirect m ty i cv -> + (forall i', + ident_ext_extends i i' -> + IdentExtMap.get i' m = IdentExtMap.get i' m') -> + cval_match_indirect m' ty i cv + with cval_match_struct_eq m m' flds i csv : + cval_match_struct m flds i csv -> + (forall i', + ident_ext_extends i i' -> + IdentExtMap.get i' m = IdentExtMap.get i' m') -> + cval_match_struct m' flds i csv + with cval_match_array_eq m m' i ty cav n : + cval_match_array m ty i cav n -> + (forall i', + ident_ext_extends i i' -> + IdentExtMap.get i' m = IdentExtMap.get i' m') -> + cval_match_array m' ty i cav n + with cval_match_hash_int_eq m m' i ty chv : + cval_match_hash_int m ty i chv -> + (forall i', + ident_ext_extends i i' -> + IdentExtMap.get i' m = IdentExtMap.get i' m') -> + cval_match_hash_int m' ty i chv. + Proof. + { intros match_indirect load_eq. + destruct match_indirect; subst. + - rewrite load_eq in H0 by (apply IIE_refl). + apply CVMIval with v; eauto. + - apply CVMIstruct, cval_match_struct_eq with m; assumption. + - apply CVMIarray, cval_match_array_eq with m; assumption. + - apply CVMIhash, cval_match_hash_int_eq with m; assumption. + } + { intros match_struct load_eq. + destruct match_struct as [ ? ? ? match_indir ]. + constructor; try assumption. + intros ? ? ? offset_eq. + destruct (match_indir _ _ _ offset_eq) + as (cv' & cv'_eq & cv'_match). + exists cv'; split. + - assumption. + - apply cval_match_indirect_eq with m. + + apply cv'_match. + + intros i' Hextends. + apply load_eq. + apply ident_ext_extends_inversion_Field in Hextends. + exact Hextends. + } + { intros match_array load_eq. + destruct match_array as [ ? ? ? ? match_indirect ]. + constructor. + intros idx idx_range. + apply cval_match_indirect_eq with m. + + apply match_indirect; assumption. + + intros i' Hextends. + apply load_eq. + apply ident_ext_extends_inversion_Index in Hextends. + exact Hextends. + } + { intros match_hash_int load_eq. + destruct match_hash_int as [ ? ? ? match_indirect ]. + constructor. + intros idx. + apply cval_match_indirect_eq with m. + + apply match_indirect; assumption. + + intros i' Hextends. + apply load_eq. + apply ident_ext_extends_inversion_Index in Hextends. + exact Hextends. + } +Qed. + diff --git a/src/core/HyperType.v b/src/core/HyperType.v new file mode 100755 index 0000000..5bc9fe7 --- /dev/null +++ b/src/core/HyperType.v @@ -0,0 +1,701 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** CompCert Clight [type] expressions bundled with Coq [Type]s and + type class declarations describing properties of the bundles. *) + +(* Standard library modules *) +Require Import BinInt. (* Z_scope, (_ | _) *) +Require Import Coq.Lists.List. +Require Import BinInt. (* Z_scope, (_ | _) *) (* Actually this changed in 8.7, need to fix. *) + +(* CompCert modules *) +Require Import cclib.Coqlib. +Require Import cclib.Errors. +(*Require Import compcert.common.Events.*) +Require Import backend.Values.HighValues. +(*Require Import compcert.common.Memory.*) +Require Import cclib.Integers. +Require Import cclib.Maps. +Require Import backend.AST. +Require Import backend.Cop. +Require Import backend.Ctypes. +Require Import backend.MemoryModel. +Require Import backend.MachineModel. +Require Import backend.Environments.ExtEnv. + +(* DeepSpec modules *) +Require Import DeepSpec.core.Cval. +Require Import DeepSpec.core.MemoryModel. +Require Import DeepSpec.lib.OProp. +Require Import DeepSpec.lib.Monad.ContOptionMonad. + +Inductive type_marker {T} (t : T) := create_type_marker : type_marker t. + +Section HYPER_TYPE. + (** * Basic Clight-Coq type linking *) + + (** A [type_pair] encapsulates [ty : type], a Clight type from [Ctypes], + and [ft : Type], the functional correspondence. *) + Inductive type_pair : Type := + Tpair : Type -> type -> type_pair. + + Definition unpair_ty tp := match tp with Tpair _ ty => ty end. + Definition unpair_ft tp := match tp with Tpair ft _ => ft end. + + + (** [HyperType] and its implementation counterpart [HyperTypeImpl] describes + the minimum requirement for a pair of types to be bundled together. + Since [HyperTypeImpl] is a super class of [HyperType] that provides no + reasoning capacity, you always want to use [HyperType] with the + generalizing binders. + + [HyperTypeImpl] contains the following four fields. + + - [ht_ty_cond] identifies C values that are valid for this [type_pair], + - [ht_ft_cond] identifies Coq values that are valid for this [type_pair], + + - [ht_implement] converts the functional value to C value, and + - [ht_spec] converts the C value to it's functional meaning + (specification). + + For example, we may bundle [tint] as defined above (which is actually + [unsigned int]) with [bool] type. We can define [ht_ty_cond] to accept + only [Vtrue] and [Vfalse], and [ht_ft_cond] accepts all possible values + of [bool] type: [true] and [false]. Naturally, [ht_implement] can be + defined ot turn [true] into [Vtrue] and [false] into [Vfalse], and + [ht_spec] does the inverse. + + To ensure that the conditions and conversions are well behaved, + [HyperType] enforces six lemmas. + + - FIXME: ht_val_has_type, ht_sem_cast_ident + - [ht_implement_returns] says that [ht_implement] always brings valid + functional value to valid C value, + - [ht_spec_returns] ensures the same for [ht_spec], + - [ht_proof_left] and [ht_proof_right] mean that [ht_spec] is + [ht_implement]'s left and right inverse, respectively. + + For the [tint] and [bool] bundle, these four trivially hold. However, + if we accidentally swapped the mapping in, say, [ht_spec], the two + [_returns] lemmas still hold, but not the two [ht_proof]. *) + (* Provides some operators on this type_pair *) + Class HyperTypeImpl (tp : type_pair) : Type := { + (* Coq value to [cval] *) + ht_cval : unpair_ft tp -> cval; + (* Default coq value *) + ht_default : unpair_ft tp; + + (* Defines what are valid coq values *) + ht_ft_cond : unpair_ft tp -> Prop; + (* defines what are valid cval *) + ht_ty_cond : cval -> Prop + := fun cv => exists f, ht_ft_cond f /\ ht_cval f = cv; + + (* Defines what can be represented by a meaningful C value. Sometimes, + * satisfying ht_ft_cond does not necessary implies ht_valid_ft_cond. take + * ATInfo from memory management as an example: ATValid corresponds to a valid + * struct data, while ATUndef means this memory region is not initialized. + * Both branches are valid coq value, but ATUndef can not be represented by + * a meaningful C value, and ht_cval usually converts it into CVany. + * This property is only required for: function args and ret, data to be stored, + * and struct to be accessed. + *) + ht_valid_ft_cond : unpair_ft tp -> Prop; + ht_valid_ty_cond : cval -> Prop + := fun cv => exists f, ht_valid_ft_cond f /\ ht_ft_cond f /\ ht_cval f = cv; + + (* Almost the same as [ht_valid_ft_cond], but uses [oprop] *) + ht_valid_ft_ocond : OProp1 (unpair_ft tp) + }. + (* Properties of [HyperTypeImpl]'s operators *) + Class HyperType (tp : type_pair)`{hti : HyperTypeImpl tp} : Prop := { + (* Valid coq value can be converted into [cval] *) + (* This is trivial since ht_cval is total, it is just a utility to help prove ht_ft_rel. *) + ht_ft_rel_core : forall f, ht_ft_cond f -> exists cv, ht_cval f = cv; + + (* The above converted [cval] is valid *) + ht_ft_rel : forall f, ht_ft_cond f -> exists cv, ht_ty_cond cv /\ ht_cval f = cv + := fun f fc => match ht_ft_rel_core f fc with ex_intro cv r => + ex_intro _ cv (conj (ex_intro _ f (conj fc r)) r) end; + (* Valid [cval] has valid preimage coq value *) + ht_ty_rel : forall cv, ht_ty_cond cv -> exists f, ht_ft_cond f /\ ht_cval f = cv + := fun cv vc => vc; + + (* Default coq value is valid *) + ht_default_ft_cond : ht_ft_cond ht_default; + + (* [ht_valid_ft_cond] is equivalent to [ht_valid_ft_ocond] *) + ht_valid_ft_ocond_same : + forall f, ht_valid_ft_cond f <-> oProp1 ht_valid_ft_ocond f + }. + + (* [ht_cval] on option types *) + Inductive ht_option_cval `{HyperTypeImpl} : + option (unpair_ft tp) -> option cval -> Prop := + | ht_none_cval : ht_option_cval None None + | ht_some_cval : forall f v, ht_cval f = v -> ht_option_cval (Some f) (Some v). + Definition ht_cval_some `{HyperTypeImpl} f vo := ht_option_cval (Some f) vo. + + (* Connect Coq value with val. *) + (* Coq value --[ht_cval]--> cval --[cval_match]--> val *) + Definition ht_rel `{HyperTypeImpl} f v := + cval_match v (ht_cval f). + Inductive ht_option_rel `{HyperTypeImpl} : + option (unpair_ft tp) -> option val -> Prop := + | ht_none_rel : ht_option_rel None None + | ht_some_rel : forall f v, ht_rel f v -> ht_option_rel (Some f) (Some v). + Definition ht_rel_some `{HyperTypeImpl} f vo := + ht_option_rel (Some f) vo. + + Class HyperByValueType (tp : type_pair) `{hti : HyperTypeImpl tp} : Prop := { + ht_by_value_access_mode : is_immediate_type (unpair_ty tp) = true; + }. + + (** By default, [ht_ty_cond], which is a class member, only takes a [val] + explicit argument, everything else is searched by [eauto] using the + typeclass instance hint database. This results in problems when multiple + [HyperType]s are defined in the environment, so we force the type-pair + argument [tp] be explicit. *) + Global Arguments ht_ty_cond tp {_} v. + Global Arguments ht_valid_ty_cond tp {_} cv. + + (** Bundling the [HyperType] instances with the [type_pair] for places + where generalizing binders are not available (in positive positions, + for example). *) + Class hyper_type_pair : Type := mk_hyper_type_pair { + _tp_type_pair : type_pair; + _tp_ty := unpair_ty _tp_type_pair; + _tp_ft := unpair_ft _tp_type_pair; + tp_hyper_type_impl :> HyperTypeImpl _tp_type_pair + }. + Definition tp_type_pair := @_tp_type_pair. + Definition tp_ty := @_tp_ty. + Definition tp_ft := @_tp_ft. + (* begin hide *) + Global Arguments mk_hyper_type_pair _ {_}. + (* end hide *) + + (** * Operations *) + (** This part encodes the interface of [Cop] in type classes. All + operations equip with: + - extra preconditions, + - the computation itself in Coq functions, + - lemma requiring the result of the computation stays in the domain, + and + - lemmas stating that the computation agrees with their counterparts + in [Cop]: + [_fcorrect] means the computation correctly casts a functional + value w.r.t. Clight semantics, and [_ccorrect] means that when the + Clight semantics exist, the function computation reproduces them. + *) + + Context `{LayerSpec : LayerSpecClass}. + + (** [HyperBuiltin] exposes the Ethereum EVM builtin operations. *) + Class HyperBuiltin0Impl (tp : type_pair) : Type := { + Hquery0 : machine_env GetHighData -> unpair_ft tp; + Hbuiltin0 : MachineModel.builtin0 + }. + + Class HyperBuiltin0 (tp : type_pair) + `{hti : HyperTypeImpl tp, !HyperBuiltin0Impl tp} : Prop := { + Hbuiltin0_returns : forall me, + ht_ft_cond (Hquery0 me); + Hbuiltin0_correct : forall me, + ht_cval (Hquery0 me) = (CVval (me_query me (Qcall0 Hbuiltin0))); + }. + + Class HyperBuiltin1Impl (arg_tp tp : type_pair) : Type := { + Hquery1 : machine_env GetHighData -> unpair_ft arg_tp -> unpair_ft tp; + Hbuiltin1 : MachineModel.builtin1 + }. + + Class HyperBuiltin1 (arg_tp tp : type_pair) + `{arg_hti : HyperTypeImpl arg_tp, hti : HyperTypeImpl tp, !HyperBuiltin1Impl arg_tp tp} : Prop := { + Hbuiltin1_returns : forall me f, + ht_ft_cond f -> + ht_ft_cond (Hquery1 me f); + Hbuiltin1_correct : forall me f v, + ht_ft_cond f -> + ht_cval f = v -> + exists v', + v = CVval v' /\ + ht_cval (Hquery1 me f) = CVval (me_query me (Qcall1 Hbuiltin1 v')); + }. + + (** [HyperUnaryOp] encodes [sem_unary_operation]. However, unlike + [sem_unary_operation] which only takes _one_ [type], these + type classes takes _two_ [type_pair]s since in some situations, + although the C type remains the same, the meaning changes and + a different functional type would fit better. For example, + when taking the boolean negation ([!]) on a [tint] associated + with [Z], the resulting C type is still [tint] but it only + makes sense to interprete the functional result in [bool]. *) + Class HyperUnaryImpl (op : unary_operation)(tp tpo : type_pair) : Type := { + (* operator pre-condition *) + Hunary_cond : unpair_ft tp -> Prop; + Hunary_ocond : OProp1 (unpair_ft tp); + + (* operator body *) + Hunary : unpair_ft tp -> unpair_ft tpo; + }. + Class HyperUnaryOp (op : unary_operation)(tp tpo : type_pair) + `{hti : HyperTypeImpl tp, htio : HyperTypeImpl tpo, !HyperUnaryImpl op tp tpo} + : Prop := { + Hunary_ocond_same : forall f, Hunary_cond f <-> oProp1 Hunary_ocond f; + (* operator post-condition *) + Hunary_returns : forall f, ht_ft_cond f -> Hunary_cond f -> + ht_ft_cond (Hunary f); + + Hunary_correct : forall f v, ht_ft_cond f -> Hunary_cond f -> ht_cval f = v -> + ht_cval_some (Hunary f) (cval_unary_operation op v (unpair_ty tp)); + }. + Class HyperUnaryPassthrough op (tp tpo : type_pair) + `{hti : HyperTypeImpl tp, htio : HyperTypeImpl tpo, !HyperUnaryImpl op tp tpo} + : Prop := { + (* Hunary_inject : forall j a b, ht_inject j a b -> ht_inject j (Hunary a) (Hunary b) *) + }. + + (** Finally, [HyperBinaryOp] encodes [sem_binary_operation]. Every thing + is standard except for [Hbinary_mem_irr], memory irrelevance, requiring + that the defined binary computation does not depend on the memory + (not even the "shape" of the memory, hence possibly two different + abstract data. *) + Class HyperBinaryImpl (op : binary_operation)(tpl tpr tpo : type_pair) + : Type := { + (* pre-condition and body *) + Hbinary_cond : unpair_ft tpl -> unpair_ft tpr -> Prop; + Hbinary_ocond : OProp2 (unpair_ft tpl) (unpair_ft tpr); + Hbinary : unpair_ft tpl -> unpair_ft tpr -> unpair_ft tpo; + }. + Class HyperBinaryOp (op : binary_operation)(tpl tpr tpo : type_pair) + `{htil : HyperTypeImpl tpl, htir : HyperTypeImpl tpr, + htio : HyperTypeImpl tpo, !HyperBinaryImpl op tpl tpr tpo} : Prop := { + Hbinary_ocond_same : forall f f', Hbinary_cond f f' <-> + oProp2 Hbinary_ocond f f'; + + (* post-condition *) + Hbinary_returns : forall f f', ht_ft_cond f -> ht_ft_cond f' -> + Hbinary_cond f f' -> ht_ft_cond (Hbinary f f'); + + Hbinary_correct : forall f f' v v', + ht_ft_cond f -> ht_ft_cond f' -> Hbinary_cond f f' -> + ht_cval f = v -> ht_cval f' = v' -> + ht_cval_some (Hbinary f f') (cval_binary_operation op v (unpair_ty tpl) + v' (unpair_ty tpr)) + + }. + Class HyperBinaryPassthrough op (tpl tpr tpo : type_pair) + `{htil : HyperTypeImpl tpl, htir : HyperTypeImpl tpr, + htio : HyperTypeImpl tpo, !HyperBinaryImpl op tpl tpr tpo} : Prop := { + }. +End HYPER_TYPE. + +Add Printing Constructor hyper_type_pair. + +(* Heterogeneous List *) +Section HLIST. + Inductive HList (F : hyper_type_pair -> Type) : + list hyper_type_pair -> Type := + | HNil : HList F nil + | HCons : forall (x : hyper_type_pair)(ls : list hyper_type_pair), + F x -> HList F ls -> HList F (x :: ls). + Global Arguments HNil {_}. + Global Arguments HCons {_} _ _ _ _. + + Definition hlist_hd {a b F} (hlist : HList F (a :: b)) : F a := + match hlist in HList _ x return match x with + | nil => unit + | cons a _ => F a + end with + | HNil => tt + | HCons _ _ x _ => x + end. + + Definition hlist_tl {a b F} (hlist : HList F (a :: b)) : HList F b := + match hlist in HList _ x return match x with + | nil => unit + | cons _ ls => HList F ls + end with + | HNil => tt + | HCons _ _ _ x => x + end. + + Fixpoint HList_map_nodep {a F A} + (f : forall htp : hyper_type_pair, F htp -> A)(hlist : HList F a) : + list A := + match hlist with + | HNil => nil + | HCons x _ y hl => cons (f x y) (HList_map_nodep f hl) + end. + + Fixpoint HList_map {a F G} + (f : forall htp : hyper_type_pair, F htp -> G htp)(hlist : HList F a) : + HList G a := + match hlist with + | HNil => HNil + | HCons x _ y hl => HCons _ _ (f x y) (HList_map f hl) + end. + + Fixpoint HList_fold_right_nodep {a F A} + (f : forall htp : hyper_type_pair, F htp -> A -> A)(a0 : A) + (hlist : HList F a) : A := + match hlist with + | HNil => a0 + | HCons x _ y hl => f x y (HList_fold_right_nodep f a0 hl) + end. + + Fixpoint list_curried_type (params : list hyper_type_pair) T : Type := + match params with + | nil => T + | htp :: l => tp_ft htp -> list_curried_type l T + end. + + Fixpoint apply_curried_func {params T} : + list_curried_type params T -> HList tp_ft params -> T := + match params with + | nil => fun t _ => t + | htp :: l => fun f hlist => @apply_curried_func l T + (f (hlist_hd hlist)) (hlist_tl hlist) + end. +End HLIST. + +Section HLIST_PROPS. + Context`{LayerSpecClass}. + + Inductive ht_list_rel : forall ls, HList tp_ft ls -> list val -> Prop := + | ht_nil_rel : ht_list_rel nil HNil nil + | ht_cons_rel : forall htp ls f fs v vs, + ht_rel f v -> ht_list_rel ls fs vs -> + ht_list_rel (cons htp ls) (HCons htp ls f fs) (cons v vs). + + Inductive ht_list_ft_cond : + forall {ls}(es1 : HList tp_ft ls), Prop := + | ht_nil_ft_cond : ht_list_ft_cond HNil + | ht_cons_ft_cond : forall htp ls e es, + ht_ft_cond e -> ht_list_ft_cond es -> + ht_list_ft_cond (HCons htp ls e es). + + Inductive ht_list_valid_ft_cond: + forall {ls} (es1 : HList tp_ft ls), Prop := + | ht_nil_valid_ft_cond: ht_list_valid_ft_cond HNil + | ht_cons_valid_ft_cond: forall htp ls e es, + ht_valid_ft_cond e -> ht_list_valid_ft_cond es -> + ht_list_valid_ft_cond (HCons htp ls e es). + + End HLIST_PROPS. + +Section HYPER_CONSTRUCTOR. + Fixpoint compose_fieldlist fld_ids fld_tps {struct fld_tps} := + match fld_tps, fld_ids with + | nil, nil => Some Fnil + | htp :: res_tps, fld_id :: res_ids => + match compose_fieldlist res_ids res_tps with + | Some res => Some (Fcons fld_id (tp_ty htp) res) + | None => None + end + | _, _ => None + end. + + Lemma compose_fieldlist_length_eq_ids_fieldlist : + forall fld_ids fld_tps fieldlist, + compose_fieldlist fld_ids fld_tps = Some fieldlist -> + length fld_ids = fieldlist_length fieldlist. + Proof. + induction fld_ids; intros. + - destruct fld_tps; try discriminate. + injection H as <-. + reflexivity. + - destruct fld_tps; try discriminate. + simpl in H. + destruct (compose_fieldlist fld_ids fld_tps) eqn:H0; try discriminate. + pose proof (IHfld_ids _ _ H0). + injection H as <-. + simpl. f_equal; assumption. + Qed. + + Definition compose_Tstruct struct_id fld_ids fld_tps := + match compose_fieldlist fld_ids fld_tps with + | Some fieldlist => Some (Tstruct struct_id fieldlist) + | None => None + end. + + Fixpoint compose_cstruct_val map fld_ids {fld_tps} + (fs : HList tp_ft fld_tps){struct fs} := + match fs with + | HNil => map + | HCons htp _ f res => match fld_ids with + | nil => map (* Impossible, but let's not bother *) + | fld_id :: res_ids => + compose_cstruct_val (PTree.set fld_id (ht_cval f) map) res_ids res + end + end. + Definition compose_cval_struct fld_ids {fld_tps}(fs : HList tp_ft fld_tps) := + CVstruct (CSmap (compose_cstruct_val (PTree.empty cval) fld_ids fs)). + + Class HyperConstructor tp `{hti : HyperTypeImpl tp} fld_ids fld_tps constr + : Prop := { + hc_compose_ty_eq : exists struct_id, + compose_Tstruct struct_id fld_ids fld_tps = Some (unpair_ty tp); + hc_fld_ids_norepet : list_norepet fld_ids; + + hc_constr_correct : forall args : HList tp_ft fld_tps, + compose_cval_struct fld_ids args + = ht_cval (apply_curried_func constr args); + hc_constr_returns : forall args : HList tp_ft fld_tps, + ht_list_ft_cond args -> + ht_ft_cond (apply_curried_func constr args) + }. + + Class HyperConstructorPassthrough + tp `{hti : HyperTypeImpl tp} fld_tps (constr: list_curried_type fld_tps (unpair_ft tp)) : Prop := { + }. + +End HYPER_CONSTRUCTOR. + +Section HYPER_LTYPE. + Context`{HM : HyperMem}. + + Definition GetOpt S V := S -> ContOpt V. + Definition SetOpt S V := V -> S -> ContOpt S. + Definition getter_of_data tp := GetOpt GetHighData (unpair_ft tp). + Definition setter_of_data tp := SetOpt GetHighData (unpair_ft tp). + + Record ltype_pair tp := { + ltype_tp_marker : type_marker tp; (* Hack to be able to use record syntax {| |}. Otherwise there would be no way to provide the value of tp to Coq's type checker. *) + + ltype_get : GetHighData -> unpair_ft tp; (* Lens with the abstract data type *) + ltype_set : unpair_ft tp -> GetHighData -> GetHighData; + + ltype_set_ocond : OProp1 GetHighData; (* These verification conditions ensure we do not read "bad" data from memory. *) + ltype_get_extra_ocond : OProp1 GetHighData; (* E.g. in user datatypes, some branches are marked "default", but we do not want to read/write those. *) + + ltype_ghost : bool; (* Should we synthesize C code for accessing this field or not? *) + ltype_ident : ident_ext; (* If so, then this is the corresponding memory ident *) + }. + Global Arguments ltype_get {_} _ _. + Global Arguments ltype_set {_} _ _ _. + Global Arguments ltype_set_ocond {_} _. + Global Arguments ltype_get_extra_ocond {_} _. + Global Arguments ltype_ghost {_} _. + Global Arguments ltype_ident {_} _. + + Class HyperLTypeDir {tp}`{HyperTypeImpl tp} (i : ident_ext) + (getter : GetHighData -> unpair_ft tp) + (setter : unpair_ft tp -> GetHighData -> GetHighData) + (setter_ocond getter_extra_ocond : OProp1 GetHighData) : Prop := { + + (* If the memory m matches the abstract data d, then reading out a part of d will match a part of m. See "extended C values" in the dissertation. *) + ltype_get_match : forall j d m, match_AbData d m j -> + oProp1 (setter_ocond /\ getter_extra_ocond)%oprop1 d -> + ht_ft_cond (getter d) /\ + cval_match_indirect m (unpair_ty tp) i (ht_cval (getter d)); + (* similar. *) + ltype_set_match : forall f j d m, ht_ft_cond f -> + mem_match j d m -> oProp1 setter_ocond d -> + forall m', (* [m'] represent post-Mem.store (potentially partial) [m] *) + (forall i', + ident_ext_extends i i' \/ + IdentExtMap.get i' (fst m') = IdentExtMap.get i' (fst m)) -> + (snd m) = (snd m') -> + cval_match_indirect (fst m') (unpair_ty tp) i (ht_cval f) -> + mem_match j (setter f d) m' + }. + + (* Those validity properties that don't mention the memory. *) + Class HyperLTypeStatic {tp}(lt : ltype_pair tp) (new_glbl : list ident) : Prop := { + }. + + (* This class combines both the "direct" and the static correctness conditions. *) + Class HyperLType {tp}`{hti : HyperTypeImpl tp}(lt : ltype_pair tp) : Prop := { + + hyper_ltype_direct : + lt.(ltype_ghost) = false -> + HyperLTypeDir lt.(ltype_ident) + lt.(ltype_get) lt.(ltype_set) + lt.(ltype_set_ocond) lt.(ltype_get_extra_ocond); + }. + + Lemma cval_match_indirect_by_value_load `{hbvt : HyperByValueType} m i cv: + cval_match_indirect m (unpair_ty tp) i cv -> + exists v, IdentExtMap.get i m = EEVal v + /\ cval_match v cv. + Proof. + assert (by_ref := ht_by_value_access_mode). + inversion 1; subst. + - exists v. + split; auto; assumption. + - rewrite <- H0 in by_ref; inversion by_ref. (* discriminate. *) + - rewrite <- H0 in by_ref; inversion by_ref. (* discriminate. *) + - rewrite <- H0 in by_ref; inversion by_ref. (* discriminate. *) + Qed. + + Lemma cval_match_indirect_by_value_store `{hbvt : HyperByValueType} m i v cv : + IdentExtMap.get i m = EEVal v -> + cval_match v cv -> + cval_match_indirect m (unpair_ty tp) i cv. + Proof. + assert (by_ref := ht_by_value_access_mode). + intros. + eapply CVMIval; eassumption. + Qed. + + (* The ltype_pair lenses always start from the full abstract datatype. + The following lenses represent a "smaller" mapping, from a struct to one of its fields (HyperFieldImpl) + or from an array to one of its entries (HypeIndexImpl). + + In HyperTypeInst.v, we combine these with ltype_pair to derive smaller ltype_pair's. + *) + Class HyperFieldImpl (tp_struct tp_field : type_pair)(id : ident) : Type := { + Hfield_get : unpair_ft tp_struct -> unpair_ft tp_field; + Hfield_set : unpair_ft tp_field -> unpair_ft tp_struct -> unpair_ft tp_struct + }. + Class HyperField (tp_struct tp_field : type_pair)(id : ident) + `{htis : HyperTypeImpl tp_struct, htif : HyperTypeImpl tp_field, + !HyperFieldImpl tp_struct tp_field id} : Prop := { + Hfield_size_pos : sizeof_words (unpair_ty tp_field) > 0; + + Hfield_get_returns : forall f, ht_ft_cond f -> ht_ft_cond (Hfield_get f); + Hfield_set_returns : forall f st, ht_ft_cond f -> ht_ft_cond st -> + ht_ft_cond (Hfield_set f st); + + Hfield_get_correct : forall st, ht_ft_cond st -> ht_valid_ft_cond st -> + ht_cval_some (Hfield_get st) + (cval_field_access id (ht_cval st) (unpair_ty tp_struct)); + + Hfield_set_correct : forall f st, ht_ft_cond f -> ht_ft_cond st -> + ht_cval_some (Hfield_set f st) + (cval_field_update id (ht_cval st) (unpair_ty tp_struct) (ht_cval f)); + + Hfield_delta_correct : + exists x ofs fList, + (unpair_ty tp_struct = Tstruct x fList /\ + struct_field fList id = Some (ofs, unpair_ty tp_field)) + }. + + (* This is used as a precondition to some lemmas about passthrough methods. *) + Class HyperFieldPassthrough (tp_struct tp_field : type_pair)(id : ident) + `{htis : HyperTypeImpl tp_struct, htif : HyperTypeImpl tp_field, + !HyperFieldImpl tp_struct tp_field id} : Prop := { + }. + + Class HyperIndexImpl (tp_array tp_elem : type_pair) : Type := { + Hindex_size : Z; + Hindex_get : Z -> unpair_ft tp_array -> unpair_ft tp_elem; + Hindex_set : Z -> unpair_ft tp_elem -> unpair_ft tp_array -> unpair_ft tp_array + }. + Class HyperIndex (tp_array tp_elem : type_pair) + `{htis : HyperTypeImpl tp_array, htif : HyperTypeImpl tp_elem, + !HyperIndexImpl tp_array tp_elem} : Prop := { + Hindex_get_returns : forall a i, (0 <= i < Hindex_size)%Z -> + ht_ft_cond a -> ht_ft_cond (Hindex_get i a); + Hindex_set_returns : forall f a i, (0 <= i < Hindex_size)%Z -> + ht_ft_cond f -> ht_ft_cond a -> ht_ft_cond (Hindex_set i f a); + Hindex_get_correct : forall a i, (0 <= i < Hindex_size)%Z -> + ht_cval_some (Hindex_get i a) + (cval_array_indexing i (ht_cval a) (unpair_ty tp_array)); + Hindex_set_correct : forall f a i, + (0 <= i < Hindex_size)%Z -> + ht_cval_some (Hindex_set i f a) + (cval_array_update i (ht_cval a) (unpair_ty tp_array) (ht_cval f)); + Hindex_array_type : + unpair_ty tp_array = Tarray (unpair_ty tp_elem) Hindex_size ; + Hindex_size_bound : 0 <= Hindex_size < Int256.modulus + }. + Class HyperIndexPassthrough (tp_array tp_elem : type_pair) + `{htis : HyperTypeImpl tp_array, htif : HyperTypeImpl tp_elem, + !HyperIndexImpl tp_array tp_elem} : Prop := { + }. + + Class HyperIntHashImpl (tp_array tp_elem : type_pair) : Type := { + Hhash_get : int256 -> unpair_ft tp_array -> unpair_ft tp_elem; + Hhash_set : int256 -> unpair_ft tp_elem -> unpair_ft tp_array -> unpair_ft tp_array + }. + + Class HyperIntHash (tp_array tp_elem : type_pair) + `{htis : HyperTypeImpl tp_array, htif : HyperTypeImpl tp_elem, + !HyperIntHashImpl tp_array tp_elem} : Prop := { + Hhash_get_returns : forall a i, + ht_ft_cond a -> ht_ft_cond (Hhash_get i a); + Hhash_set_returns : forall f a i, + ht_ft_cond f -> ht_ft_cond a -> ht_ft_cond (Hhash_set i f a); + Hhash_get_correct : forall a i, + ht_cval_some (Hhash_get i a) + (cval_hashmap_lookup i (ht_cval a) (unpair_ty tp_array)); + Hhash_set_correct : forall f a i, + ht_cval_some (Hhash_set i f a) + (cval_hashmap_update i (ht_cval a) (unpair_ty tp_array) (ht_cval f)); + Hhash_type : + unpair_ty tp_array = Thashmap tint (unpair_ty tp_elem) + }. + + +End HYPER_LTYPE. + +(* This is "legacy code": this definition turned out to not be quite right, but a lot of things were implemented using it, + so we provide wrappers to convert to the newer version above. *) +Section ISO_TYPE. + Class HyperTypeIsoImpl (tp : type_pair) : Type := { + ht_iso_ft_cond : unpair_ft tp -> Prop; + ht_iso_ty_cond : val -> Prop; + + ht_iso_default : unpair_ft tp; + + ht_implement : unpair_ft tp -> val; + ht_spec : val -> unpair_ft tp + }. + Class HyperTypeIso (tp : type_pair)`{hti : HyperTypeIsoImpl tp} : Prop := { + ht_implement_returns : forall f, ht_iso_ft_cond f -> ht_iso_ty_cond (ht_implement f); + ht_spec_returns : forall v, ht_iso_ty_cond v -> ht_iso_ft_cond (ht_spec v); + + ht_iso_default_ft_cond : ht_iso_ft_cond ht_iso_default; + + ht_proof_left : forall f, ht_iso_ft_cond f -> ht_spec (ht_implement f) = f; + ht_proof_right : forall v, ht_iso_ty_cond v -> ht_implement (ht_spec v) = v + }. + Global Arguments ht_iso_ty_cond tp {_} v. + + Context`{HyperTypeIso}. + + Global Instance hyper_type_iso_impl : HyperTypeImpl tp := { + ht_cval f := CVval (ht_implement f); + ht_ft_cond := ht_iso_ft_cond; + ht_default := ht_iso_default; + ht_valid_ft_cond f := True; + ht_valid_ft_ocond := otrue1; + (* ht_inject j a b := a = b *) + }. + Global Instance hyper_type_iso : HyperType tp. + Proof. esplit. + - (* ht_ft_rel_core *) + intros f fc. + eexists; reflexivity. + - (* ht_default_ft_cond *) + simpl. + apply ht_iso_default_ft_cond. + - (* ht_valid_ft_ocond_same *) + simpl; split; trivial. + Qed. +End ISO_TYPE. diff --git a/src/core/HyperTypeInst.v b/src/core/HyperTypeInst.v new file mode 100755 index 0000000..20e1af8 --- /dev/null +++ b/src/core/HyperTypeInst.v @@ -0,0 +1,2412 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** Primitive type pairing and operators on them. *) + +(* Standard library modules *) +Require Import ZArith. +Require Import Znumtheory. + +(* CompCert modules *) +Require Import backend.Values.HighValues. +Require Import cclib.Integers. +Require Import backend.AST. +Require Import backend.Cop. +Require Import backend.Ctypes. + +(* DeepSpec modules *) + +Require Import DeepSpec.core.Cval. +Require Import DeepSpec.core.HyperType. +Require Import DeepSpec.core.MemoryModel. +Require Import DeepSpec.lib.OProp. +Require Import DeepSpec.lib.Monad.ContOptionMonad. + +(** * [unit] corresponds to [void] *) +Notation tvoid_unit := (Tpair unit Tvoid). + +Section HYPER_TYPE_UNIT. + Global Instance void_unit_impl : HyperTypeImpl tvoid_unit := { + ht_cval := fun _ => CVany; + ht_ft_cond := fun _ => True; + ht_default := tt; + ht_valid_ft_cond := fun _ => False; + ht_valid_ft_ocond := ofalse1; + (* ht_inject j a b := True *) + }. + Global Instance void_unit : HyperType tvoid_unit. + Proof. esplit. + - (* ht_ft_rel_core *) + intros. + exists CVany. + reflexivity. + + - (* ht_default_ft_cond *) + exact I. + + - (* ht_valid_ft_ocond_same *) + simpl; split; trivial. + Qed. + + Definition void_unit_pair := mk_hyper_type_pair tvoid_unit. + +End HYPER_TYPE_UNIT. + +Section HYPER_LTYPE_COMPOSITION. + Context`{HM : HyperMem}. + Context {tp : type_pair}. + Variable (whole : ltype_pair tp). + + Context`{hti : HyperTypeImpl tp, hlt : !HyperLType whole}. + + Section HYPER_LTYPE_FIELD. + Context {tp_field : type_pair}. + Variable (tid : ident). + Context`{htif : !HyperTypeImpl tp_field, + hfieldi : !HyperFieldImpl tp tp_field tid, + hfield : !HyperField tp tp_field tid}. + + Definition field_ltype_pair : ltype_pair tp_field := {| + ltype_tp_marker := create_type_marker tp_field; + + ltype_get s := Hfield_get (whole.(ltype_get) s); + ltype_set v s := + whole.(ltype_set) + (Hfield_set v (whole.(ltype_get) s)) + s; + + ltype_set_ocond := + (whole.(ltype_set_ocond) /\ whole.(ltype_get_extra_ocond))%oprop1; + ltype_get_extra_ocond := + omap1 (fun p y => p (whole.(ltype_get) y)) ht_valid_ft_ocond; + + ltype_ghost := whole.(ltype_ghost); + ltype_ident := (Field (whole.(ltype_ident)) tid) + |}. + + Lemma nat_le_discrete (n m : nat) : + (n <= m -> exists n', n' + n = m)%nat. + Proof. + induction 1. + - exists 0%nat; reflexivity. + - destruct IHle as [ n'' <- ]. + exists (S n''). + reflexivity. + Qed. + + Context {htw : HyperType tp}. + Global Instance field_ltype : HyperLType field_ltype_pair. + Proof. + esplit. + esplit. + - (* ltype_get_match *) + intros j d m mm dc. + + simpl in dc. + rewrite oand1_distr in dc; destruct dc as [ whole_dc dc ]. + rewrite OProp1map1, <- ht_valid_ft_ocond_same in dc; try exact I. + + destruct hlt as [hlt_dir]. + simpl in H. specialize (hlt_dir H). + assert (IH := ltype_get_match (HyperLTypeDir := hlt_dir) _ _ _ mm whole_dc). + + assert (field_match := Hfield_get_correct (whole.(ltype_get) d) + (proj1 IH) dc). + inversion field_match as [| f v ? f_eq access_some ]; + subst; clear field_match. + + (* Unify information from [HyperField] and [Cval] *) + destruct (cval_sem_field_access _ _ _ _ (eq_sym access_some)) + as (strid' & flds' & t2 & delta & ty_eq' & field_offset_eq' & match_imply). + + destruct (Hfield_delta_correct) as [strid [ofs [flds [ty_eq field_offset_eq]]]]. + rewrite ty_eq' in ty_eq; clear ty_eq'. + inversion ty_eq; subst. + rewrite field_offset_eq' in field_offset_eq; clear field_offset_eq'. + inversion field_offset_eq; subst. + + split. + + (* ht_ft_cond *) + apply Hfield_get_returns, IH. + + (* cval_match_indirect *) + unfold ltype_ident; simpl. + apply match_imply, IH. + + - (* ltype_set_match *) + simpl. + intros f j d m fc mm dc m' disjoint_eq cm. + admit. + Admitted. + End HYPER_LTYPE_FIELD. + + Section HYPER_LTYPE_INDEX. + Context {tp_elem : type_pair}. + Context`{htif : HyperTypeImpl tp_elem, + hidxi : !HyperIndexImpl tp tp_elem, + hidx : !HyperIndex tp tp_elem}. + + Variable (idx : Z). + + Definition indexing_ltype_pair : ltype_pair tp_elem := {| + ltype_tp_marker := create_type_marker tp_elem; + + ltype_get s := Hindex_get idx (whole.(ltype_get) s); + ltype_set v s := + whole.(ltype_set) + (Hindex_set idx v (whole.(ltype_get) s)) + s; + + ltype_set_ocond := + (whole.(ltype_set_ocond) /\ whole.(ltype_get_extra_ocond))%oprop1; + ltype_get_extra_ocond := otrue1; + + ltype_ghost := whole.(ltype_ghost); + ltype_ident := (Index whole.(ltype_ident) (Int256.repr idx)) + |}. + + Require Import cclib.Coqlib. + + Global Instance indexing_ltype : + forall idx_in_bound : 0 <= idx < Hindex_size, + HyperLType indexing_ltype_pair. + Proof. + (* Prepare commonly used assertions *) + assert (array_type_eq := Hindex_array_type). + + intros idx_in_bound. + + + esplit. + esplit. + + - (* ltype_get_match *) + intros j d m mm dc. + + simpl in dc. + rewrite oand1_distr in dc. + apply proj1 in dc. + + assert (index_match := Hindex_get_correct (whole.(ltype_get) d) + _ idx_in_bound). + inversion index_match as [| f v ? f_eq index_some ]; + subst; clear index_match. + + (* Unify information from [HyperIndex] and [Cval] *) + destruct (cval_sem_array_indexing _ _ _ _ (eq_sym index_some)) + as (t2 & n & ty_eq' & match_imply). + rewrite array_type_eq in ty_eq'. + inversion ty_eq'; subst; clear ty_eq'. + + destruct hlt as [hlt_dir]. + specialize (hlt_dir H). + assert (IH := ltype_get_match (HyperLTypeDir := hlt_dir) _ _ _ mm dc). + split. + + (* ht_ft_cond *) + apply Hindex_get_returns, IH. + apply idx_in_bound. + + (* cval_match_indirect *) + simpl. + apply match_imply, IH. + + - (* ltype_set_match *) + simpl. + intros f j d m fc mm dc m' disjoint_eq same_d cm. + destruct hlt as [hlt_dir]. + specialize (hlt_dir H). + apply (ltype_set_match (HyperLTypeDir:=hlt_dir)) with (m:=m); try assumption. + + (* ht_ft_cond (Hindex_set f (ltype_get whole d)) *) + apply Hindex_set_returns; try assumption. + eapply ltype_get_match; [apply mm | apply dc]. + + (* oProp1 whole.(ltype_set_ocond) d *) + rewrite oand1_distr in dc; apply dc. + + (* disjoint_eq *) + intros i'. + destruct (disjoint_eq i') as [disjoint_eq' | disjoint_eq' ]. + * left. + apply ident_ext_extends_inversion_Index in disjoint_eq'. + exact disjoint_eq'. + * right. + exact disjoint_eq'. + + (* cval_match_indirect *) + assert (index_get_rel := Hindex_set_correct f (ltype_get whole d) + _ idx_in_bound). + inversion index_get_rel as [| ? ? v_eq f_eq index_get_eq ]; + subst; clear index_get_rel. + unfold cval_array_update in index_get_eq. + assert (orig := proj2 (ltype_get_match _ _ _ (proj2 mm) dc)). + destruct (ht_cval _) as [| | | | []] in index_get_eq at 2, orig; + try discriminate index_get_eq. + rewrite array_type_eq in index_get_eq, orig |- *. + destruct c as [map]. + destruct (zle 0 idx && zlt idx Hindex_size); + try discriminate index_get_eq. + injection index_get_eq as ->. + + inversion orig as [ ? ? ? ? ? | | ? ? ? ? orig_array | ]; + try match goal with [ H: is_immediate_type _ = true |- _] => now inversion H end; subst. + inversion orig_array as [ ? ? ? ? orig_indir ]; subst. + constructor; constructor. + intros idx' idx_in_range'. + specialize (orig_indir _ idx_in_range'). + destruct (zeq idx' idx) as [ -> | idx_ne ]. + * rewrite Maps.ZMap.gss. + apply cm. + * { rewrite Maps.ZMap.gso by assumption. + apply cval_match_indirect_eq with (fst m). + - apply orig_indir. + - intros i' Hextends. + destruct (disjoint_eq i') as [disjoint_eq' | disjoint_eq']. + contradict disjoint_eq'. + eapply ident_ext_extends_disjoint_Index with (o1:= Int256.repr idx'). + + intros Heq. + destruct Hindex_size_bound. + apply Int256.repr_injective in Heq; omega. + + assumption. + + symmetry; assumption. + } + Qed. + End HYPER_LTYPE_INDEX. + + Section HYPER_LTYPE_HASH. + Context {tp_elem : type_pair}. + Context`{htif : HyperTypeImpl tp_elem, + hidxi : !HyperIntHashImpl tp tp_elem, + hidx : !HyperIntHash tp tp_elem}. + + Variable (idx : int256). + + Definition inthash_ltype_pair : ltype_pair tp_elem := {| + ltype_tp_marker := create_type_marker tp_elem; + + ltype_get s := Hhash_get idx (whole.(ltype_get) s); + ltype_set v s := + whole.(ltype_set) + (Hhash_set idx v (whole.(ltype_get) s)) + s; + + ltype_set_ocond := + (whole.(ltype_set_ocond) /\ whole.(ltype_get_extra_ocond))%oprop1; + ltype_get_extra_ocond := otrue1; + + ltype_ghost := whole.(ltype_ghost); + ltype_ident := (Index whole.(ltype_ident) idx) + |}. + + + Global Instance inthash_ltype : + HyperLType inthash_ltype_pair. + Proof. + assert (hash_type_eq := Hhash_type). + esplit. + esplit. + - (* ltype_get_match *) + intros j d m mm dc. + + simpl in dc. + rewrite oand1_distr in dc. + apply proj1 in dc. + + assert (hash_match := Hhash_get_correct (whole.(ltype_get) d) idx). + inversion hash_match as [| f v ? f_eq hash_some ]; + subst; clear hash_match. + + (* Unify information from [HyperIndex] and [Cval] *) + destruct (cval_sem_hashmap_lookup _ _ _ _ (eq_sym hash_some)) + as (t2 & ty_eq' & match_imply). + rewrite hash_type_eq in ty_eq'. + inversion ty_eq'; subst; clear ty_eq'. + + destruct hlt as [hlt_dir]. + specialize (hlt_dir H). + assert (IH := ltype_get_match (HyperLTypeDir := hlt_dir) _ _ _ mm dc). + split. + + (* ht_ft_cond *) + apply Hhash_get_returns, IH. + + (* cval_match_indirect *) + simpl. + apply match_imply, IH. + + - (* ltype_set_match *) + simpl. + intros f j d m fc mm dc m' disjoint_eq same_d cm. + destruct hlt as [hlt_dir]. + specialize (hlt_dir H). + apply (ltype_set_match (HyperLTypeDir:=hlt_dir)) with (m:=m); try assumption. + + (* ht_ft_cond (Hindex_set f (ltype_get whole d)) *) + apply Hhash_set_returns; try assumption. + eapply ltype_get_match; [apply mm | apply dc]. + + (* oProp1 whole.(ltype_set_ocond) d *) + rewrite oand1_distr in dc; apply dc. + + (* disjoint_eq *) + intros i'. + destruct (disjoint_eq i') as [disjoint_eq' | disjoint_eq' ]. + * left. + apply ident_ext_extends_inversion_Index in disjoint_eq'. + exact disjoint_eq'. + * right. + exact disjoint_eq'. + + + (* cval_match_indirect *) + assert (index_get_rel := Hhash_set_correct f (ltype_get whole d) + idx). + inversion index_get_rel as [| ? ? v_eq f_eq index_get_eq ]; + subst; clear index_get_rel. + unfold cval_hashmap_update in index_get_eq. + assert (orig := proj2 (ltype_get_match _ _ _ (proj2 mm) dc)). + destruct (ht_cval _) as [| | | | []] in index_get_eq at 2, orig; + try discriminate index_get_eq. + rewrite hash_type_eq in index_get_eq, orig |- *. + injection index_get_eq as ->. + inversion orig as [ ? ? ? ? ? | | | ? ? ? orig_array ]; + try match goal with [ H: is_immediate_type _ = true |- _] => now inversion H end; subst. + inversion orig_array as [ ? ? ? orig_indir ]; subst. + constructor; constructor. + intros idx'. + specialize (orig_indir idx'). + destruct (Int256.eq_dec idx' idx) as [ -> | idx_ne ]. + * rewrite Maps.Int256Map.gss. + apply cm. + * { rewrite Maps.Int256Map.gso by assumption. + apply cval_match_indirect_eq with (fst m). + - apply orig_indir. + - intros i' Hextends. + destruct (disjoint_eq i') as [disjoint_eq' | disjoint_eq']. + contradict disjoint_eq'. + eapply ident_ext_extends_disjoint_Index with (o1:= idx'). + assumption. + assumption. + symmetry; assumption. } + Qed. + + End HYPER_LTYPE_HASH. + +End HYPER_LTYPE_COMPOSITION. + +(** * Integers as booleans *) +Notation tint_bool := (Tpair bool tint). + +Section Integer256Extra. +Lemma modulus_gt_zero : Int256.modulus > 0. +Proof. + unfold Z.gt, Z.compare, Int256.modulus, two_power_nat; + reflexivity. +Qed. + +Theorem sub_sub : forall x y z, + Int256.sub (Int256.sub x y) z = Int256.sub x (Int256.add y z). +Proof. + intros x y z. + rewrite (Int256.sub_add_opp x), Int256.sub_add_l. + symmetry; apply Int256.sub_add_r. +Qed. + +Theorem add_repr : forall x y, + Int256.add (Int256.repr x) (Int256.repr y) = Int256.repr (x + y). +Proof. + intros x y. + rewrite Int256.add_unsigned. + apply Int256.eqm_samerepr, Int256.eqm_add; + rewrite Int256.unsigned_repr_eq; + apply Int256.eqmod_sym, Int256.eqmod_mod; + reflexivity. +Qed. + +Theorem sub_repr : forall x y, + Int256.sub (Int256.repr x) (Int256.repr y) = Int256.repr (x - y). +Proof. + intros x y. + unfold Int256.sub. + apply Int256.eqm_samerepr, Int256.eqm_sub; + rewrite Int256.unsigned_repr_eq; + apply Int256.eqmod_sym, Int256.eqmod_mod; + reflexivity. +Qed. + +Theorem mul_repr : forall x y, + Int256.mul (Int256.repr x) (Int256.repr y) = Int256.repr (x * y). +Proof. + intros x y. + unfold Int256.mul. + apply Int256.eqm_samerepr, Int256.eqm_mult; + rewrite Int256.unsigned_repr_eq; + apply Int256.eqmod_sym, Int256.eqmod_mod; + reflexivity. +Qed. + +Theorem add_shifted : forall x y z, + Int256.add (Int256.sub x z) (Int256.add y z) = Int256.add x y. +Proof. + intros x y z. + rewrite Int256.add_permut, Int256.sub_add_opp, Int256.add_assoc, + (Int256.add_commut (Int256.neg z)), Int256.add_neg_zero, Int256.add_zero. + apply Int256.add_commut. +Qed. +End Integer256Extra. + +Section HYPER_TYPE_BOOL. + Local Open Scope Z_scope. + + Lemma small_modulo : 0 mod Int256.modulus = 0 /\ 1 mod Int256.modulus = 1. + Proof. + split; apply Zmod_small; split; + try solve + [ apply Zle_refl | + unfold Int256.modulus, two_power_nat, Z.lt, Z.compare; trivial ]. + + apply Zlt_le_weak. + assert (1 = Z.succ 0). + unfold Z.succ, Z.add; reflexivity. + rewrite H; apply Zle_lt_succ. + apply Zle_refl. + Qed. + Definition zero_mod_modulus := proj1 small_modulo. + Definition one_mod_modulus := proj2 small_modulo. + Ltac rewrite_unsigned_repr := + try unfold Int256.zero, Int256.one; + try rewrite Int256.unsigned_repr_eq; + try rewrite zero_mod_modulus; + try rewrite one_mod_modulus. + + Local Instance int_bool_iso_impl : HyperTypeIsoImpl tint_bool := { + ht_iso_ty_cond v := match v with + | Vint i => Int256.unsigned i = 0 \/ Int256.unsigned i = 1 + | _ => False + end; + ht_iso_ft_cond f := True; + + ht_iso_default := false; + + ht_implement f := Vint (if f then Int256.one else Int256.zero); + ht_spec v := match v with + | Vint i => Int256.unsigned i =? 1 + | _ => false + end + }. + Local Instance int_bool_iso : HyperTypeIso tint_bool. + Proof. esplit. + - (* ht_implement_returns *) + intros; simpl. + destruct f; simpl; rewrite_unsigned_repr; [ right | left ]; reflexivity. + + - (* ht_spec_returns *) + intros; simpl. + trivial. + + - (* ht_iso_default_ft_cond *) + exact I. + + - (* ht_proof_left *) + intros; simpl. + destruct f; simpl; rewrite_unsigned_repr; reflexivity. + + - (* ht_proof_right *) + intros v vc; simpl in vc |- *. + destruct v; try contradiction. + destruct vc as [ H | H ]; rewrite H; simpl; + unfold Vtrue, Vfalse, Int256.zero, Int256.one; + rewrite <- H; rewrite (Int256.repr_unsigned i); reflexivity. + Qed. + Global Instance int_bool_impl : HyperTypeImpl tint_bool := _. + Global Instance int_bool : HyperType tint_bool := _. + + Definition int_bool_pair := + @mk_hyper_type_pair tint_bool int_bool_impl. + + Lemma ht_ty_cond_0_1 : + ht_ty_cond tint_bool (CVval (Vint (Int256.repr 0))) /\ + ht_ty_cond tint_bool (CVval (Vint (Int256.repr 1))). + Proof. + split; [ exists false | exists true ]; split; + solve [ exact I | reflexivity ]. + Qed. + Definition int_bool_zero := proj1 ht_ty_cond_0_1. + Definition int_bool_one := proj2 ht_ty_cond_0_1. + + Global Instance int_bool_notbool_impl + : HyperUnaryImpl Onotbool tint_bool tint_bool := { + Hunary_cond ft := True; + Hunary_ocond := otrue1; + Hunary := negb + }. + Global Instance int_bool_notbool : HyperUnaryOp Onotbool tint_bool tint_bool. + Proof. esplit. + - (* Hunary_ocond_same *) + reflexivity. + + - (* Hunary_returns *) + simpl; trivial. + + - (* Hunary_correct *) + intros f v fc oc rel. + simpl in rel; rewrite <- rel. + destruct f; constructor; reflexivity. + Qed. + Global Instance int_bool_notbool_passthrough + : HyperUnaryPassthrough Onotbool tint_bool tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Global Instance int_bool_or_impl + : HyperBinaryImpl Oor tint_bool tint_bool tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := orb + }. + + Global Instance int_bool_or + : HyperBinaryOp Oor tint_bool tint_bool tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel'; rewrite <- rel, <- rel'. + destruct f; destruct f'; + constructor; reflexivity. + Qed. + Global Instance int_bool_or_passthrough + : HyperBinaryPassthrough Oor tint_bool tint_bool tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Global Instance int_bool_xor_impl + : HyperBinaryImpl Oxor tint_bool tint_bool tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := xorb + }. + Global Instance int_bool_xor + : HyperBinaryOp Oxor tint_bool tint_bool tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel'; rewrite <- rel, <- rel'. + destruct f; destruct f'; + constructor; reflexivity. + Qed. + Global Instance int_bool_xor_passthrough + : HyperBinaryPassthrough Oxor tint_bool tint_bool tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Global Instance int_bool_and_impl + : HyperBinaryImpl Oand tint_bool tint_bool tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := andb + }. + Global Instance int_bool_and + : HyperBinaryOp Oand tint_bool tint_bool tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel'; rewrite <- rel, <- rel'. + destruct f; destruct f'; + constructor; reflexivity. + Qed. + Global Instance int_bool_and_passthrough + : HyperBinaryPassthrough Oand tint_bool tint_bool tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Global Instance int_bool_eq_impl + : HyperBinaryImpl Oeq tint_bool tint_bool tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary f f' := negb (xorb f f') + }. + Global Instance int_bool_eq + : HyperBinaryOp Oeq tint_bool tint_bool tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel'; rewrite <- rel, <- rel'. + destruct f; destruct f'; + constructor; reflexivity. + Qed. + Global Instance int_bool_eq_passthrough + : HyperBinaryPassthrough Oeq tint_bool tint_bool tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Global Instance int_bool_ne_impl + : HyperBinaryImpl One tint_bool tint_bool tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := xorb + }. + Global Instance int_bool_ne + : HyperBinaryOp One tint_bool tint_bool tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel'; rewrite <- rel, <- rel'. + destruct f; destruct f'; + constructor; reflexivity. + Qed. + Global Instance int_bool_ne_passthrough + : HyperBinaryPassthrough One tint_bool tint_bool tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + (* Other operations, e.g. addition, are not defined on tint_bool so + there are no other instances, as demonstrated by the lacking of + ``HyperBinaryOp Oadd tint_bool ...'' *) +End HYPER_TYPE_BOOL. + +(** * Integers as Z *) +Section HYPER_TYPE_INT. + Local Open Scope Z_scope. + + Class IntegerBound (bound : Z) : Prop := { + integer_bound_within_modulus : Int256.modulus >= bound; + integer_bound_positive : 0 < bound + }. + + Definition Z_bounded (bound : Z){_ : IntegerBound bound} := Z. + + Typeclasses Opaque Z_bounded. + + Opaque Int256.repr. + +Section BOUNDED. + (* All the definitions will need these from here on. *) + Variable bound : Z. + Context `{bound_cond : IntegerBound bound}. + + Definition tint_Z_bounded + := Tpair (Z_bounded bound) tint. + + Lemma unsigned_repr_eq (f : Z_bounded bound)(cond : -1 < f /\ bound > f) + : Int256.unsigned (Int256.repr f) = f. + Proof. + rewrite Int256.unsigned_repr_eq. + apply Zmod_small. + destruct cond as [ m1_lt_f bound_gt_f ]. + split. + - assert (0 = Z.succ (-1)). + reflexivity. + rewrite H; apply Zlt_le_succ; assumption. + - apply Zgt_lt. + eapply Zle_gt_trans; + [ apply Zge_le; apply integer_bound_within_modulus | exact bound_gt_f ]. + Qed. + + Lemma bounded_repr_injective : forall (f f' : Z_bounded bound) + (cond : -1 < f /\ bound > f) + (cond' : -1 < f' /\ bound > f'), + (Int256.repr f) = (Int256.repr f') -> f = f'. + Proof. + intros. + pose integer_bound_within_modulus. + apply Int256.repr_injective in H; auto; omega. + Qed. + + Instance int_Z_iso_impl : HyperTypeIsoImpl tint_Z_bounded := { + ht_iso_ty_cond v := match v with + | Vint i => bound > Int256.unsigned i + | _ => False + end; + ht_iso_ft_cond f := -1 < f /\ bound > f; + + ht_iso_default := 0; + + ht_implement f := (Vint (Int256.repr f)); + ht_spec v := match v with + | Vint i => Int256.unsigned i + | _ => 0 + end + }. + + + Instance int_Z_iso : HyperTypeIso tint_Z_bounded. + Proof. esplit. + - (* ht_implement_returns *) + intros f fc; simpl in fc |- *. + rewrite unsigned_repr_eq; + apply fc. + + - (* ht_spec_returns *) + intros v vc; simpl in vc |- *. + destruct v; try contradiction. + split; try assumption. + destruct i. + unfold Int256.unsigned, Int256.intval. + apply intrange. + + - (* ht_iso_default_ft_cond *) + split; [ reflexivity |]. + apply Z.lt_gt, integer_bound_positive. + + - (* ht_proof_left *) + intros f fc; simpl in fc |- *; unfold Z_bounded. + apply unsigned_repr_eq, fc. + + - (* ht_proof_right *) + intros v vc; simpl in vc |- *. + destruct v; try contradiction. + rewrite (Int256.repr_unsigned i). + reflexivity. + Qed. + Instance int_Z_impl : HyperTypeImpl tint_Z_bounded := _. + Instance int_Z : HyperType tint_Z_bounded := _. + + (* Notbool is not defined on tint_Z_bounded so the instance is not + defined. Even if it was defined, the never satisfied + pre-condition would render it useless in practice. *) + + Instance int_Z_bounded_add_impl + : HyperBinaryImpl Oadd tint_Z_bounded tint_Z_bounded tint_Z_bounded := { + Hbinary_cond f f' := f + f' < bound; + Hbinary_ocond := oprop2 (fun f f' => f + f' < bound); + Hbinary := Z.add + }. + Instance int_Z_bounded_add + : HyperBinaryOp Oadd tint_Z_bounded tint_Z_bounded tint_Z_bounded := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros f f' fc fc' oc. + split. + + (* -1 < Hbinary f f' *) + apply Z.le_succ_l; change (Z.succ (-1)) with (Z.succ (-1) + Z.succ (-1)). + apply Z.add_le_mono; apply Z.le_succ_l; [ apply fc | apply fc' ]. + + (* bound > Hbinary f f' *) + apply Z.lt_gt, oc. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel' |- *. + rewrite <- rel, <- rel'. + constructor. + rewrite add_repr; reflexivity. + Qed. + Instance int_Z_bounded_add_passthrough + : HyperBinaryPassthrough Oadd tint_Z_bounded tint_Z_bounded tint_Z_bounded. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Instance int_Z_bounded_sub_impl + : HyperBinaryImpl Osub tint_Z_bounded tint_Z_bounded tint_Z_bounded := { + Hbinary_cond f f' := f >= f'; + Hbinary_ocond := oprop2 (fun f f' => f >= f'); + Hbinary := Z.sub + }. + Instance int_Z_bounded_sub + : HyperBinaryOp Osub tint_Z_bounded tint_Z_bounded tint_Z_bounded := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros f f' fc fc' oc. + split. + + (* -1 < Hbinary f f' *) + simpl. + rewrite <- Z.le_succ_l, <- Z.le_add_le_sub_r. + apply Z.ge_le, oc. + + (* bound > Hbinary f f' *) + simpl in fc, fc' |- *. + omega. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel' |- *. + rewrite <- rel, <- rel'. + constructor. + simpl. + rewrite sub_repr; reflexivity. + Qed. + Instance int_Z_bounded_sub_passthrough + : HyperBinaryPassthrough Osub tint_Z_bounded tint_Z_bounded tint_Z_bounded. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Instance int_Z_bounded_mul_impl + : HyperBinaryImpl Omul tint_Z_bounded tint_Z_bounded tint_Z_bounded := { + Hbinary_cond f f' := f * f' < bound; + Hbinary_ocond := oprop2 (fun f f' => f * f' < bound); + Hbinary := Z.mul + }. + Instance int_Z_bounded_mul + : HyperBinaryOp Omul tint_Z_bounded tint_Z_bounded tint_Z_bounded := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros f f' fc fc' oc. + split. + + (* -1 < Hbinary f f' *) + apply Z.le_succ_l, Z.mul_nonneg_nonneg; + change 0 with (Z.succ (-1)); apply Z.le_succ_l; + [ apply fc | apply fc' ]. + + (* bound > Hbinary f f' *) + apply Z.lt_gt, oc. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel' |- *. + rewrite <- rel, <- rel'. + constructor. + rewrite mul_repr; reflexivity. + Qed. + Instance int_Z_bounded_mul_passthrough + : HyperBinaryPassthrough Omul tint_Z_bounded tint_Z_bounded tint_Z_bounded. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Instance int_Z_bounded_mod_impl + : HyperBinaryImpl Omod tint_Z_bounded tint_Z_bounded tint_Z_bounded := { + Hbinary_cond f f' := f' <> 0; + Hbinary_ocond := oprop2 (fun f f' => f' <> 0); + Hbinary := Z.modulo + }. + Instance int_Z_bounded_mod + : HyperBinaryOp Omod tint_Z_bounded tint_Z_bounded tint_Z_bounded := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros f f' fc fc' oc. + simpl in oc, fc' |- *. + split. + + (* -1 < Hbinary f f' *) + rewrite <- Z.le_succ_l. + apply Z_mod_lt. + apply proj1 in fc'. + rewrite <- Z.le_succ_l, Z.lt_eq_cases in fc'. + destruct fc' as [ lt | eq ]; + [ apply Z.lt_gt, lt | contradiction (oc (eq_sym eq)) ]. + + (* bound > Hbinary f f' *) + eapply Zgt_trans; [ apply fc' |]. + apply Z.lt_gt, Z_mod_lt. + apply proj1 in fc'. + rewrite <- Z.le_succ_l, Z.lt_eq_cases in fc'. + destruct fc' as [ lt | eq ]; + [ apply Z.lt_gt, lt | contradiction (oc (eq_sym eq)) ]. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel', fc, fc', oc |- *. + + assert (f_eq := unsigned_repr_eq _ fc). + assert (f_eq' := unsigned_repr_eq _ fc'). + + rewrite <- rel, <- rel'. + simpl. (* unfold sem_mod, sem_binarith; simpl. *) + assert (f_zero_spec := Int256.eq_spec (Int256.repr f') Int256.zero). + destruct (Int256.eq (Int256.repr f') Int256.zero). + + apply (f_equal Int256.unsigned) in f_zero_spec. + rewrite Int256.unsigned_zero, f_eq' in f_zero_spec. + contradiction (oc f_zero_spec). + + constructor. + unfold Int256.modu; rewrite f_eq, f_eq'. + reflexivity. + Qed. + Instance int_Z_bounded_mod_passthrough + : HyperBinaryPassthrough Omod tint_Z_bounded tint_Z_bounded tint_Z_bounded. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + (* The div operation is not defined in the backend yet. *) + Instance int_Z_bounded_div_impl + : HyperBinaryImpl Odiv tint_Z_bounded tint_Z_bounded tint_Z_bounded := { + Hbinary_cond f f' := f' <> 0; + Hbinary_ocond := oprop2 (fun f f' => f' <> 0); + Hbinary := Z.div + }. + Instance int_Z_bounded_div + : HyperBinaryOp Odiv tint_Z_bounded tint_Z_bounded tint_Z_bounded := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros f f' fc fc' oc. + simpl in oc, fc, fc' |- *. + + assert (f_pos' : 0 < f'). + { apply proj1 in fc'. + rewrite <- Z.le_succ_l, Z.lt_eq_cases in fc'. + destruct fc' as [ lt | eq ]; + [ exact lt | contradiction (oc (eq_sym eq)) ]. + } + + split. + + (* -1 < Hbinary f f' *) + rewrite <- Z.le_succ_l. + apply Z_div_pos; [ apply Z.lt_gt, f_pos' |]. + rewrite <- Z.le_succ_l in fc. + apply fc. + + (* bound > Hbinary f f' *) + apply Z.lt_gt, Zdiv_lt_upper_bound; [ exact f_pos' |]. + eapply Z.lt_le_trans; [ apply Z.gt_lt, fc |]. + replace bound with (bound * 1) at 1 by apply Z.mul_1_r. + apply Zmult_le_compat_l; + [| apply Z.lt_le_incl, integer_bound_positive ]. + rewrite <- Z.le_succ_l in f_pos'. + exact f_pos'. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel', fc, fc', oc |- *. + + assert (f_eq := unsigned_repr_eq _ fc). + assert (f_eq' := unsigned_repr_eq _ fc'). + + rewrite <- rel, <- rel'. + simpl. + assert (f_zero_spec := Int256.eq_spec (Int256.repr f') Int256.zero). + destruct (Int256.eq (Int256.repr f') Int256.zero). + + apply (f_equal Int256.unsigned) in f_zero_spec. + rewrite Int256.unsigned_zero, f_eq' in f_zero_spec. + contradiction (oc f_zero_spec). + + constructor. + simpl. + unfold Int256.divu; rewrite f_eq, f_eq'. + reflexivity. + Qed. + Instance int_Z_bounded_div_passthrough + : HyperBinaryPassthrough Odiv tint_Z_bounded tint_Z_bounded tint_Z_bounded. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Instance int_Z_bounded_eq_impl + : HyperBinaryImpl Oeq tint_Z_bounded tint_Z_bounded tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Z.eqb + }. + Instance int_Z_bounded_eq + : HyperBinaryOp Oeq tint_Z_bounded tint_Z_bounded tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; exact I. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel' |- *. + rewrite <- rel, <- rel'. + constructor. + destruct (val_eq_dec (Vint (Int256.repr f)) (Vint (Int256.repr f'))). + + pose (bounded_repr_injective f f' fc fc'). + rewrite e0; try congruence. + rewrite Z.eqb_refl. + reflexivity. + + rewrite (proj2 (Z.eqb_neq f f')). + * reflexivity. + * intros Heq. + subst. + apply n. + reflexivity. + Qed. + Instance int_Z_bounded_eq_passthrough + : HyperBinaryPassthrough Oeq tint_Z_bounded tint_Z_bounded tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Instance int_Z_bounded_ne_impl + : HyperBinaryImpl One tint_Z_bounded tint_Z_bounded tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary f f' := negb (Z.eqb f f') + }. + Instance int_Z_bounded_ne + : HyperBinaryOp One tint_Z_bounded tint_Z_bounded tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; exact I. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel' |- *. + rewrite <- rel, <- rel'. + constructor. + destruct (val_eq_dec (Vint (Int256.repr f)) (Vint (Int256.repr f'))). + + pose (bounded_repr_injective f f' fc fc'). + rewrite e0; try congruence. + rewrite Z.eqb_refl. + reflexivity. + + rewrite (proj2 (Z.eqb_neq f f')). + * reflexivity. + * intros Heq. + subst. + apply n. + reflexivity. + Qed. + Instance int_Z_bounded_ne_passthrough + : HyperBinaryPassthrough One tint_Z_bounded tint_Z_bounded tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Instance int_Z_bounded_lt_impl + : HyperBinaryImpl Olt tint_Z_bounded tint_Z_bounded tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Z.ltb + }. + Instance int_Z_bounded_lt + : HyperBinaryOp Olt tint_Z_bounded tint_Z_bounded tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; exact I. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel' |- *. + rewrite <- rel, <- rel'. + constructor. + unfold Int256.cmpu, Int256.ltu, Coqlib.zlt. + rewrite 2 unsigned_repr_eq; try assumption. + destruct (Z_lt_dec f f') as [ dec_lt | dec_ge ], + (f ? f') eqn:ltb_eq. + + reflexivity. + + exfalso. + rewrite Z.gtb_ltb, (proj2 (Z.ltb_lt _ _) dec_lt) in ltb_eq; discriminate. + + contradict dec_ge. + rewrite Z.gtb_ltb in ltb_eq. + apply Z.ltb_lt, ltb_eq. + + reflexivity. + Qed. + Instance int_Z_bounded_gt_passthrough + : HyperBinaryPassthrough Ogt tint_Z_bounded tint_Z_bounded tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Instance int_Z_bounded_le_impl + : HyperBinaryImpl Ole tint_Z_bounded tint_Z_bounded tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Z.leb + }. + Instance int_Z_bounded_le + : HyperBinaryOp Ole tint_Z_bounded tint_Z_bounded tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; exact I. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel' |- *. + rewrite <- rel, <- rel'. + constructor. + unfold Int256.cmpu, Int256.ltu, Coqlib.zlt. + rewrite 2 unsigned_repr_eq; try assumption. + destruct (Z_lt_dec f' f) as [ dec_lt | dec_ge ], + (f <=? f') eqn:ltb_eq. + + exfalso. + rewrite Z.leb_antisym, (proj2 (Z.ltb_lt _ _) dec_lt) in ltb_eq; + discriminate. + + reflexivity. + + reflexivity. + + contradict dec_ge. + apply (f_equal negb) in ltb_eq. + rewrite Z.leb_antisym, negb_involutive in ltb_eq. + apply Z.ltb_lt, ltb_eq. + Qed. + Instance int_Z_bounded_le_passthrough + : HyperBinaryPassthrough Ole tint_Z_bounded tint_Z_bounded tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Instance int_Z_bounded_ge_impl + : HyperBinaryImpl Oge tint_Z_bounded tint_Z_bounded tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Z.geb + }. + Instance int_Z_bounded_ge + : HyperBinaryOp Oge tint_Z_bounded tint_Z_bounded tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; exact I. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel' |- *. + rewrite <- rel, <- rel'. + constructor. + unfold Int256.cmpu, Int256.ltu, Coqlib.zlt. + rewrite 2 unsigned_repr_eq; try assumption. + destruct (Z_lt_dec f f') as [ dec_lt | dec_ge ], + (f >=? f') eqn:geb_eq. + + exfalso. + apply Z.geb_le, Z.le_ge in geb_eq. + exact (geb_eq dec_lt). + + reflexivity. + + reflexivity. + + contradict dec_ge. + rewrite Z.geb_leb in geb_eq. + apply Z.leb_gt, geb_eq. + Qed. + Instance int_Z_bounded_ge_passthrough + : HyperBinaryPassthrough Oge tint_Z_bounded tint_Z_bounded tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + +End BOUNDED. + + Global Instance modulus_bound : IntegerBound Int256.modulus := { + integer_bound_within_modulus := Z.le_ge _ _ (Zle_refl Int256.modulus); + integer_bound_positive := Z.gt_lt _ _ (Int256.modulus_pos) + }. + + Definition tint_Z32 := tint_Z_bounded Int256.modulus. + Typeclasses Transparent tint_Z32 tint_Z_bounded. + + Definition int_Z32_impl := int_Z_impl Int256.modulus. + Definition int_Z32 := int_Z Int256.modulus. + (* Definition int_Z32_by_value := int_Z_by_value Int256.modulus.* *) + Definition int_Z32_pair := @mk_hyper_type_pair tint_Z32 int_Z32_impl (*int_Z32*). + (*Definition int_Z32_self_cast := int_Z_self_cast Int256.modulus. + Definition int_Z32_arg_ret := int_Z_arg_ret. *) + Existing Instance int_Z32_impl. + + Lemma int_Z32_ty_cond z + : ht_ty_cond tint_Z32 (CVval (HighValues.Vint (Int256.repr z))). + Proof. + unfold ht_ty_cond, int_Z32_impl, int_Z_impl. + exists (Int256.Z_mod_modulus z). + split. + - (* ht_ft_cond (Int256.Z_mod_modulus z) *) + split; [| apply Z.lt_gt ]; apply Int256.Z_mod_modulus_range'. + - (* ht_rel (Int256.Z_mod_modulus z) (Vint (Int256.repr z)) *) + simpl; do 2 f_equal. + rewrite Int256.Z_mod_modulus_eq. + apply Int256.eqm_samerepr, Int256.eqmod_sym, Int256.eqmod_mod, Int256.modulus_pos. + Qed. + + Definition int_Z32_add_impl := int_Z_bounded_add_impl Int256.modulus. + Definition int_Z32_sub_impl := int_Z_bounded_sub_impl Int256.modulus. + Definition int_Z32_mul_impl := int_Z_bounded_mul_impl Int256.modulus. + Definition int_Z32_mod_impl := int_Z_bounded_mod_impl Int256.modulus. + Definition int_Z32_div_impl := int_Z_bounded_div_impl Int256.modulus. + + Definition int_Z32_add := int_Z_bounded_add Int256.modulus. + Definition int_Z32_sub := int_Z_bounded_sub Int256.modulus. + Definition int_Z32_mul := int_Z_bounded_mul Int256.modulus. + Definition int_Z32_mod := int_Z_bounded_mod Int256.modulus. + Definition int_Z32_div := int_Z_bounded_div Int256.modulus. + + Definition int_Z32_add_passthrough := int_Z_bounded_add_passthrough Int256.modulus. + Definition int_Z32_sub_passthrough := int_Z_bounded_sub_passthrough Int256.modulus. + Definition int_Z32_mul_passthrough := int_Z_bounded_mul_passthrough Int256.modulus. + Definition int_Z32_mod_passthrough := int_Z_bounded_mod_passthrough Int256.modulus. + Definition int_Z32_div_passthrough := int_Z_bounded_div_passthrough Int256.modulus. + + Definition int_Z32_eq_impl := int_Z_bounded_eq_impl Int256.modulus. + Definition int_Z32_ne_impl := int_Z_bounded_ne_impl Int256.modulus. + Definition int_Z32_lt_impl := int_Z_bounded_lt_impl Int256.modulus. + Definition int_Z32_gt_impl := int_Z_bounded_gt_impl Int256.modulus. + Definition int_Z32_le_impl := int_Z_bounded_le_impl Int256.modulus. + Definition int_Z32_ge_impl := int_Z_bounded_ge_impl Int256.modulus. + + Definition int_Z32_eq := int_Z_bounded_eq Int256.modulus. + Definition int_Z32_ne := int_Z_bounded_ne Int256.modulus. + Definition int_Z32_lt := int_Z_bounded_lt Int256.modulus. + Definition int_Z32_gt := int_Z_bounded_gt Int256.modulus. + Definition int_Z32_le := int_Z_bounded_le Int256.modulus. + Definition int_Z32_ge := int_Z_bounded_ge Int256.modulus. + + Definition int_Z32_eq_passthrough := int_Z_bounded_eq_passthrough Int256.modulus. + Definition int_Z32_ne_passthrough := int_Z_bounded_ne_passthrough Int256.modulus. + Definition int_Z32_lt_passthrough := int_Z_bounded_lt_passthrough Int256.modulus. + Definition int_Z32_gt_passthrough := int_Z_bounded_gt_passthrough Int256.modulus. + Definition int_Z32_le_passthrough := int_Z_bounded_le_passthrough Int256.modulus. + Definition int_Z32_ge_passthrough := int_Z_bounded_ge_passthrough Int256.modulus. + +End HYPER_TYPE_INT. + +Notation Z32 := (Z_bounded Int256.modulus). + +Existing Instances int_Z32_impl int_Z32 (*int_Z32_by_value int_Z32_arg_ret*). +Existing Instances int_Z32_add_impl int_Z32_sub_impl int_Z32_mul_impl + int_Z32_mod_impl int_Z32_div_impl. +Existing Instances int_Z32_add int_Z32_sub int_Z32_mul int_Z32_mod int_Z32_div. +Existing Instances int_Z32_add_passthrough int_Z32_sub_passthrough + int_Z32_mul_passthrough int_Z32_mod_passthrough + int_Z32_div_passthrough. +Existing Instances int_Z32_eq_impl int_Z32_ne_impl int_Z32_lt_impl + int_Z32_gt_impl int_Z32_le_impl int_Z32_ge_impl. +Existing Instances int_Z32_eq int_Z32_ne int_Z32_lt int_Z32_gt + int_Z32_le int_Z32_ge. +Existing Instances int_Z32_eq_passthrough int_Z32_ne_passthrough + int_Z32_lt_passthrough int_Z32_gt_passthrough + int_Z32_le_passthrough int_Z32_ge_passthrough. + +Require Import Omega. + +Section HYPER_TYPE_UNSIGNED. + (* Local Open Scope Z_scope. *) + + Definition tint_U := Tpair int256 tint. + + Instance int_U_iso_impl : HyperTypeIsoImpl tint_U := + { + ht_iso_ft_cond f := True; + ht_iso_ty_cond v := match v with + | Vint i => Int256.modulus > Int256.unsigned i + | _ => False + end; + ht_iso_default := Int256.zero; + ht_implement f := (Vint f); + ht_spec v := match v with + | Vint i => i + | _ => Int256.zero + end + }. + + Instance int_U_iso : HyperTypeIso tint_U. + Proof. + esplit. + - (* ht_implement_returns *) + intros f fc. simpl in *. + destruct f. + simpl. omega. + - (* ht_spec_returns *) + intros v vc; simpl in *. + destruct v; try contradiction. + split; auto. + - (* ht_iso_default_ft_cond *) + unfold ht_iso_ft_cond. simpl; trivial. + - (* ht_proof_left *) + intros f fc; reflexivity. + - (* ht_proof_right *) + intros [] vc; simpl in *; try contradiction. + reflexivity. + Qed. + + Instance int_U_impl : HyperTypeImpl tint_U := _. + Instance int_U : HyperType tint_U := _. + + (**[int_U_add]******************************************) + Instance int_U_add_impl : HyperBinaryImpl Oadd tint_U tint_U tint_U := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary x y := Int256.add x y + }. + + Instance int_U_add : HyperBinaryOp Oadd tint_U tint_U tint_U := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + intros f f' fc fc' oc. + split; omega. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in *. subst v v'. + constructor. + reflexivity. + Qed. + Instance int_U_add_passthrough : HyperBinaryPassthrough Oadd tint_U tint_U tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_sub]******************************************) + Instance int_U_sub_impl : HyperBinaryImpl Osub tint_U tint_U tint_U := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary x y := Int256.sub x y + }. + + Instance int_U_sub : HyperBinaryOp Osub tint_U tint_U tint_U := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + intros f f' fc fc' oc. + split; omega. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in *. + subst v v'. + constructor. reflexivity. + Qed. + Instance int_U_sub_passthrough : HyperBinaryPassthrough Osub tint_U tint_U tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + + (**[int_U_mul]******************************************) + Instance int_U_mul_impl : HyperBinaryImpl Omul tint_U tint_U tint_U := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary x y := Int256.mul x y + }. + + Instance int_U_mul : HyperBinaryOp Omul tint_U tint_U tint_U := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + intros f f' fc fc' oc. + split; simpl; omega. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in *. + subst v v'. + constructor. + reflexivity. + Qed. + Instance int_U_mul_passthrough : HyperBinaryPassthrough Omul tint_U tint_U tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_mod]******************************************) + Instance int_U_mod_impl : HyperBinaryImpl Omod tint_U tint_U tint_U := + { + Hbinary_cond f f' := f' <> Int256.zero; + Hbinary_ocond := oprop2 (fun f f' => f' <> Int256.zero); + Hbinary := Int256.modu + }. + + Instance int_U_mod : HyperBinaryOp Omod tint_U tint_U tint_U := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + intros f f' fc fc' oc. + simpl in *. + split; omega. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'. simpl in *. + constructor. + reflexivity. + Qed. + Instance int_U_mod_passthrough : HyperBinaryPassthrough Omod tint_U tint_U tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_div]******************************************) + Instance int_U_div_impl : HyperBinaryImpl Odiv tint_U tint_U tint_U := + { + Hbinary_cond f f' := f' <> Int256.zero; + Hbinary_ocond := oprop2 (fun f f' => f' <> Int256.zero); + Hbinary := Int256.divu + }. + Instance int_U_div : HyperBinaryOp Odiv tint_U tint_U tint_U := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + intros f f' fc fc' oc. + simpl in *. + split; omega. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'; simpl in *. + constructor. reflexivity. + Qed. + Instance int_U_div_passthrough : HyperBinaryPassthrough Odiv tint_U tint_U tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_eq]******************************************) + Instance int_U_eq_impl : HyperBinaryImpl Oeq tint_U tint_U tint_bool := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Int256.cmpu Ceq + }. + + Opaque val_eq_dec. + + Instance int_U_eq : HyperBinaryOp Oeq tint_U tint_U tint_bool := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'; simpl in *. + constructor; simpl. + assert (eqspec := Int256.eq_spec f f'). + destruct (val_eq_dec (Vint f) (Vint f')); + destruct (Int256.eq f f'); + simpl in *; + try reflexivity; + congruence. + Qed. + Instance int_U_eq_passthrough : HyperBinaryPassthrough Oeq tint_U tint_U tint_bool . + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_ne]******************************************) + Instance int_U_ne_impl : HyperBinaryImpl One tint_U tint_U tint_bool := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Int256.cmpu Cne + }. + + Instance int_U_ne : HyperBinaryOp One tint_U tint_U tint_bool := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'; simpl in *. + constructor; simpl. + assert (eqspec := Int256.eq_spec f f'). + destruct (val_eq_dec (Vint f) (Vint f')); + destruct (Int256.eq f f'); + simpl in *; + try reflexivity; + congruence. + Qed. + Instance int_U_ne_passthrough : HyperBinaryPassthrough One tint_U tint_U tint_bool. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + + (**[int_U_lt]******************************************) + Instance int_U_lt_impl : HyperBinaryImpl Olt tint_U tint_U tint_bool := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Int256.cmpu Clt + }. + + Instance int_U_lt : HyperBinaryOp Olt tint_U tint_U tint_bool := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'. + constructor; simpl in *. + destruct (Int256.ltu f f'); reflexivity. + Qed. + Instance int_U_lt_passthrough : HyperBinaryPassthrough Olt tint_U tint_U tint_bool. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_gt]******************************************) + Instance int_U_gt_impl : HyperBinaryImpl Ogt tint_U tint_U tint_bool := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Int256.cmpu Cgt + }. + + Instance int_U_gt : HyperBinaryOp Ogt tint_U tint_U tint_bool := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'. + constructor; simpl in *. + destruct (Int256.ltu f' f); reflexivity. + Qed. + Instance int_U_gt_passthrough : HyperBinaryPassthrough Ogt tint_U tint_U tint_bool. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_le]******************************************) + Instance int_U_le_impl : HyperBinaryImpl Ole tint_U tint_U tint_bool := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Int256.cmpu Cle + }. + + Instance int_U_le : HyperBinaryOp Ole tint_U tint_U tint_bool := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'. + constructor; simpl in *. + destruct (Int256.ltu f' f); reflexivity. + Qed. + Instance int_U_le_passthrough : HyperBinaryPassthrough Ole tint_U tint_U tint_bool. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_ge]******************************************) + Instance int_U_ge_impl : HyperBinaryImpl Oge tint_U tint_U tint_bool := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Int256.cmpu Cge + }. + + Instance int_U_ge : HyperBinaryOp Oge tint_U tint_U tint_bool := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'. + constructor; simpl in *. + destruct (Int256.ltu f f'); reflexivity. + Qed. + Instance int_U_ge_passthrough : HyperBinaryPassthrough Oge tint_U tint_U tint_bool. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_notint]******************************************) + Instance int_U_notint_impl : HyperUnaryImpl Onotint tint_U tint_U := + { + Hunary_cond ft := True; + Hunary_ocond := otrue1; + Hunary := Int256.not + }. + + Instance int_U_notint : HyperUnaryOp Onotint tint_U tint_U. + Proof. + esplit. + - (* Hunary_ocond_same *) + reflexivity. + - (* Hunary_returns *) + trivial. + - (* Hunary_correct *) + simpl. + intros f v fc oc rel. + subst v. constructor. reflexivity. + Qed. + Instance int_U_notint_passthrough : HyperUnaryPassthrough Onotint tint_U tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_and]******************************************) + Instance int_U_and_impl : HyperBinaryImpl Oand tint_U tint_U tint_U := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Int256.and + }. + + Instance int_U_and : HyperBinaryOp Oand tint_U tint_U tint_U := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in *. subst v v'. + constructor. reflexivity. + Qed. + Instance int_U_and_passthrough : HyperBinaryPassthrough Oand tint_U tint_U tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + + (**[int_U_or]******************************************) + Instance int_U_or_impl : HyperBinaryImpl Oor tint_U tint_U tint_U := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Int256.or + }. + Instance int_U_or : HyperBinaryOp Oor tint_U tint_U tint_U := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in *. subst v v'. + constructor. reflexivity. + Qed. + Instance int_U_or_passthrough : HyperBinaryPassthrough Oor tint_U tint_U tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_shl]******************************************) + Instance int_U_shl_impl : HyperBinaryImpl Oshl tint_U tint_Z32 tint_U := + { + Hbinary_cond f f' := f' < Int256.zwordsize; + Hbinary_ocond := oprop2 (fun f f' => f' < Int256.zwordsize); + Hbinary x y := Int256.shl x (Int256.repr y) + }. + Lemma lt_zwordsize_lt_iwordsize : + forall f, + -1 < f < Int256.zwordsize -> + Int256.ltu (Int256.repr f) Int256.iwordsize = true. + Proof. + intros. + unfold Int256.ltu, zlt. + unfold Int256.iwordsize. + assert (Int256.wordsize = 256%nat) by reflexivity. + assert (Int256.zwordsize = 256) by reflexivity. + rewrite H1 in *. + repeat rewrite (unsigned_repr_eq _ _). + - destruct (Z_lt_dec f 256); auto. + omega. + - split. + + omega. + + unfold Int256.modulus. + rewrite H0. + rewrite two_power_nat_equiv. + reflexivity. + - unfold Int256.modulus. + rewrite two_power_nat_equiv. + rewrite H0. + split. + + omega. + + cut (2 ^ Z.of_nat 256 > 256). + * omega. + * reflexivity. + Qed. + + Instance int_U_shl : HyperBinaryOp Oshl tint_U tint_Z32 tint_U := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'; simpl in *. + constructor. + reflexivity. + Qed. + Instance int_U_shl_passthrough : HyperBinaryPassthrough Oshl tint_U tint_Z32 tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + (**[int_U_shr]******************************************) + Instance int_U_shr_impl : HyperBinaryImpl Oshr tint_U tint_Z32 tint_U := + { + Hbinary_cond f f' := f' < Int256.zwordsize; + Hbinary_ocond := oprop2 (fun f f' => f' < Int256.zwordsize); + Hbinary x y := Int256.shru x (Int256.repr y) + }. + Instance int_U_shr : HyperBinaryOp Oshr tint_U tint_Z32 tint_U := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'; simpl in *. + constructor. + reflexivity. + Qed. + Instance int_U_shr_passthrough : HyperBinaryPassthrough Oshr tint_U tint_Z32 tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + + (**[int_U_xor]******************************************) + Instance int_U_xor_impl : HyperBinaryImpl Oxor tint_U tint_U tint_U := + { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := Int256.xor + }. + Instance int_U_xor : HyperBinaryOp Oxor tint_U tint_U tint_U := + { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + - (* Hbinary_returns *) + trivial. + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in *. subst v v'. + constructor; simpl. + reflexivity. + Qed. + Instance int_U_xor_passthrough : HyperBinaryPassthrough Oxor tint_U tint_U tint_U. + (* Proof. + esplit. + simpl; congruence. + Qed. *) + + Typeclasses Transparent tint_U. + Definition int_U_pair := @mk_hyper_type_pair tint_U int_U_impl. + Existing Instance int_U_impl. + + Lemma int_U_ty_cond z : + ht_ty_cond tint_U (CVval (HighValues.Vint z)). + Proof. + unfold ht_ty_cond, int_U_impl. + exists z. + split. + - split. + - simpl; do 2 f_equal. + Qed. + +(* Definition tint_U_naturally_aligned : naturally_aligned (unpair_ty tint_U) := (eq_refl _). *) + +End HYPER_TYPE_UNSIGNED. + + + +Existing Instances int_U_impl int_U (* int_U_by_value int_U_arg_ret*). +Existing Instances + int_U_add_impl int_U_sub_impl int_U_mul_impl + int_U_mod_impl int_U_div_impl. +Existing Instances + int_U_add int_U_sub + int_U_mul int_U_mod int_U_div. +Existing Instances + int_U_add_passthrough int_U_sub_passthrough + int_U_mul_passthrough int_U_mod_passthrough int_U_div_passthrough. +Existing Instances + int_U_eq_impl int_U_ne_impl + int_U_lt_impl int_U_gt_impl int_U_le_impl int_U_ge_impl. +Existing Instances + int_U_eq int_U_ne + int_U_lt int_U_gt int_U_le int_U_ge. +Existing Instances + int_U_eq_passthrough int_U_ne_passthrough + int_U_lt_passthrough int_U_gt_passthrough + int_U_le_passthrough int_U_ge_passthrough. +Existing Instances + int_U_notint_impl int_U_and_impl int_U_or_impl + int_U_shl_impl int_U_shr_impl int_U_xor_impl. +Existing Instances + int_U_notint int_U_and int_U_or + int_U_shl int_U_shr int_U_xor. +Existing Instances + int_U_notint_passthrough int_U_and_passthrough int_U_or_passthrough + int_U_shl_passthrough int_U_shr_passthrough int_U_xor_passthrough. + +Require Import backend.MachineModel. + +Section WITH_DATA. + +Context `{LayerSpec : LayerSpecClass}. + +Local Arguments HyperBuiltin0 {_} {tp} {hti} (HyperBuiltin0Impl0). +Local Arguments HyperBuiltin1 {_} {arg_tp} {tp} {arg_hti} {hti} (HyperBuiltin1Impl0). + +Instance builtin0_address_impl : HyperBuiltin0Impl tint_U + := Build_HyperBuiltin0Impl tint_U me_address Baddress. +Instance builtin0_address : HyperBuiltin0 builtin0_address_impl. +constructor; reflexivity. +Qed. + +Instance builtin0_origin_impl : HyperBuiltin0Impl tint_U + := Build_HyperBuiltin0Impl tint_U me_origin Borigin. +Instance builtin0_origin : HyperBuiltin0 builtin0_origin_impl. +constructor; reflexivity. +Qed. + +Instance builtin0_caller_impl : HyperBuiltin0Impl tint_U + := Build_HyperBuiltin0Impl tint_U me_caller Bcaller. +Instance builtin0_caller : HyperBuiltin0 builtin0_caller_impl. +constructor; reflexivity. +Qed. + +Instance builtin0_callvalue_impl : HyperBuiltin0Impl tint_U + := Build_HyperBuiltin0Impl tint_U me_callvalue Bcallvalue. +Instance builtin0_callvalue : HyperBuiltin0 builtin0_callvalue_impl. +constructor; reflexivity. +Qed. + +Instance builtin0_coinbase_impl : HyperBuiltin0Impl tint_U + := Build_HyperBuiltin0Impl tint_U me_coinbase Bcoinbase. +Instance builtin0_coinbase : HyperBuiltin0 builtin0_coinbase_impl. +constructor; reflexivity. +Qed. + +Instance builtin0_timestamp_impl : HyperBuiltin0Impl tint_U + := Build_HyperBuiltin0Impl tint_U me_timestamp Btimestamp. +Instance builtin0_timestamp : HyperBuiltin0 builtin0_timestamp_impl. +constructor; reflexivity. +Qed. + +Instance builtin0_number_impl : HyperBuiltin0Impl tint_U + := Build_HyperBuiltin0Impl tint_U me_number Bnumber. +Instance builtin0_number : HyperBuiltin0 builtin0_number_impl. +constructor; reflexivity. +Qed. + +Instance builtin1_balance_impl : HyperBuiltin1Impl tint_U tint_U + := Build_HyperBuiltin1Impl tint_U tint_U me_balance Bbalance. +Instance builtin1_balance : HyperBuiltin1 builtin1_balance_impl. +constructor. +- (* HBuiltin1_return *) + constructor. +- (* HBuiltin1_correct *) + intros. + simpl in H0. + exists (Vint f). + split. + + symmetry. apply H0. + + reflexivity. +Qed. + +Instance builtin1_blockhash_impl : HyperBuiltin1Impl tint_U tint_U + := Build_HyperBuiltin1Impl tint_U tint_U me_blockhash Bblockhash. +Instance builtin1_blockhash : HyperBuiltin1 builtin1_blockhash_impl. +constructor. +- (* HBuiltin1_return *) + constructor. +- (* HBuiltin1_correct *) + intros. + simpl in H0. + exists (Vint f). + split. + + symmetry. apply H0. + + reflexivity. +Qed. + +End WITH_DATA. + +(* For the keccak-256 hashes, we represent them symbolically, as a + tree of the hash operations that have been done. This corresponds + to an assumption that the hash functions are injective. *) + +Require Import backend.SymbolicKeccak. + +Inductive hashvalue : Set := +| hashval_int256 : int256 -> hashvalue +| hashval_hash1 : hashvalue -> hashvalue +| hashval_hash2 : hashvalue -> hashvalue -> hashvalue. + +Fixpoint val_of_hashvalue hv := + match hv with + | hashval_int256 i => Vint i + | hashval_hash1 hv1 => sha_1 (val_of_hashvalue hv1) + | hashval_hash2 hv1 hv2 => sha_2 (val_of_hashvalue hv1) (val_of_hashvalue hv2) + end. + +Lemma val_of_hashvalue_injective : forall hv hv', + val_of_hashvalue hv = val_of_hashvalue hv' -> + hv = hv'. +Proof. + induction hv; destruct hv'; simpl; intros; + try match goal with + | [ H : Vint ?i = sha_1 ?v |- _] => pose (sha_1_range v i) + | [ H : sha_1 ?v = Vint ?i |- _] => pose (sha_1_range v i) + | [ H : Vint ?i = sha_2 ?v1 ?v2 |- _] => pose (sha_2_range v1 v2 i) + | [ H : sha_2 ?v1 ?v2 = Vint ?i |- _] => pose (sha_2_range v1 v2 i) + | [ H : sha_1 _ = sha_1 _ |- _ ] => pose (sha_1_injective _ _ H) + | [ H : sha_2 _ _ = sha_2 _ _ |- _] => destruct (sha_2_injective _ _ _ _ H) + | [ H : sha_1 ?u = sha_2 ?v1 ?v2 |- _] => pose (sha_1_sha_2_range v1 v2 u) + | [ H : sha_2 ?v1 ?v2 = sha_1 ?u |- _] => pose (sha_1_sha_2_range v1 v2 u) + end; + try congruence; f_equal; eauto. +Qed. + +Definition tint_hashvalue := Tpair hashvalue tint. + +Instance int_hashvalue_impl : HyperTypeImpl tint_hashvalue := { + ht_cval hv := CVval (val_of_hashvalue hv); + ht_default := hashval_int256 Int256.zero; + ht_ft_cond hv := True; + ht_valid_ft_cond hv := True; + ht_valid_ft_ocond := otrue1; + (* ht_inject i hv1 hv2 := hv1 = hv2 *) +}. + +Instance int_hashvalue : HyperType tint_hashvalue. + constructor. + - intros hv _. + eexists. reflexivity. + - simpl. tauto. + - simpl. tauto. +Qed. + +Definition int_hashvalue_pair := + @mk_hyper_type_pair tint_hashvalue int_hashvalue_impl. + +Global Instance int_hashvalue_hash1_impl + : HyperUnaryImpl Osha_1 tint_hashvalue tint_hashvalue := { + Hunary_cond ft := True; + Hunary_ocond := otrue1; + Hunary := hashval_hash1 +}. + +Global Instance int_hashvalue_hash1 : HyperUnaryOp Osha_1 tint_hashvalue tint_hashvalue. +Proof. + esplit. + - (* Hunary_ocond_same *) + reflexivity. + - (* Hunary_returns. *) + simpl. + trivial. + - (* Hunary_correct. *) + intros f v fc oc rel. + rewrite <- rel. + simpl. + apply ht_some_cval. + simpl. + reflexivity. +Qed. + +Global Instance int_hashvalue_hash1_passthrough : HyperUnaryPassthrough Osha_1 tint_hashvalue tint_hashvalue. +(* Proof. + constructor. + simpl. + intros; congruence. +Qed. *) + +Global Instance int_hashvalue_hash1_U_impl + : HyperUnaryImpl Osha_1 tint_U tint_hashvalue := { + Hunary_cond ft := True; + Hunary_ocond := otrue1; + Hunary f := (hashval_hash1 (hashval_int256 f)) +}. + +Global Instance int_hashvalue_hash1_U : HyperUnaryOp Osha_1 tint_U tint_hashvalue. +Proof. + esplit. + - (* Hunary_ocond_same *) + reflexivity. + - (* Hunary_returns. *) + simpl. + trivial. + - (* Hunary_correct. *) + intros f v fc oc rel. + rewrite <- rel. + simpl. + apply ht_some_cval. + simpl. + reflexivity. +Qed. + +Global Instance int_hashvalue_hash1_U_passthrough : HyperUnaryPassthrough Osha_1 tint_U tint_hashvalue. +(* Proof. + constructor. + simpl. + intros; congruence. +Qed. *) + +Global Instance int_hashvalue_hash2_impl : + HyperBinaryImpl Osha_2 tint_hashvalue tint_hashvalue tint_hashvalue := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary := hashval_hash2 + }. + + Global Instance int_hashvalue_hash2 + : HyperBinaryOp Osha_2 tint_hashvalue tint_hashvalue tint_hashvalue := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel'; rewrite <- rel, <- rel'. + simpl. + apply ht_some_cval. + simpl. + reflexivity. + Qed. + + Global Instance int_hashvalue_hash2_passthrough + : HyperBinaryPassthrough Osha_2 tint_hashvalue tint_hashvalue tint_hashvalue := { + }. + (* Proof. + simpl. + intros. + congruence. + Qed. *) + + +Global Instance int_hashvalue_hash2_U1impl : + HyperBinaryImpl Osha_2 tint_U tint_hashvalue tint_hashvalue := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary x y := hashval_hash2 (hashval_int256 x) y + }. + + Global Instance int_hashvalue_hash2_U1 + : HyperBinaryOp Osha_2 tint_U tint_hashvalue tint_hashvalue := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel'; rewrite <- rel, <- rel'. + simpl. + apply ht_some_cval. + simpl. + reflexivity. + Qed. + + Global Instance int_hashvalue_hash2_U1_passthrough + : HyperBinaryPassthrough Osha_2 tint_U tint_hashvalue tint_hashvalue := { + }. + (* Proof. + simpl. + intros. + congruence. + Qed. *) + + +Global Instance int_hashvalue_hash2_U2_impl : + HyperBinaryImpl Osha_2 tint_hashvalue tint_U tint_hashvalue := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary x y := hashval_hash2 x (hashval_int256 y) + }. + + Global Instance int_hashvalue_hash2_U2 + : HyperBinaryOp Osha_2 tint_hashvalue tint_U tint_hashvalue := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel'; rewrite <- rel, <- rel'. + simpl. + apply ht_some_cval. + simpl. + reflexivity. + Qed. + + Global Instance int_hashvalue_hash2_U2_passthrough + : HyperBinaryPassthrough Osha_2 tint_hashvalue tint_U tint_hashvalue := { + }. + (* Proof. + simpl. + intros. + congruence. + Qed. *) + +Global Instance int_hashvalue_hash2_U12_impl : + HyperBinaryImpl Osha_2 tint_U tint_U tint_hashvalue := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary x y := hashval_hash2 (hashval_int256 x) (hashval_int256 y) + }. + + Global Instance int_hashvalue_hash2_U12 + : HyperBinaryOp Osha_2 tint_U tint_U tint_hashvalue := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + simpl in rel, rel'; rewrite <- rel, <- rel'. + simpl. + apply ht_some_cval. + simpl. + reflexivity. + Qed. + + Global Instance int_hashvalue_hash2_U12_passthrough + : HyperBinaryPassthrough Osha_2 tint_U tint_U tint_hashvalue := { + }. + (* Proof. + simpl. + intros. + congruence. + Qed. *) + + Definition hashvalue_eq_dec : forall (hv hv' : hashvalue), + {hv = hv'} + {hv <> hv'}. + decide equality. + apply Int256.eq_dec. + Defined. + + Definition hashvalue_eqb (hv hv' : hashvalue) : bool := + if hashvalue_eq_dec hv hv' then true else false. + + Global Instance hashvalue_bool_eq_impl + : HyperBinaryImpl Oeq tint_hashvalue tint_hashvalue tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary f f' := hashvalue_eqb f f' + }. + + Global Instance hashvalue_bool_eq + : HyperBinaryOp Oeq tint_hashvalue tint_hashvalue tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'; simpl in *. + constructor; simpl. + unfold hashvalue_eqb. + destruct (hashvalue_eq_dec f f'); + destruct (val_eq_dec (val_of_hashvalue f) (val_of_hashvalue f')); + try reflexivity. + + subst. + congruence. + + contradict n. + auto using val_of_hashvalue_injective. + Qed. + Global Instance hashvalue_bool_eq_passthrough + : HyperBinaryPassthrough Oeq tint_hashvalue tint_hashvalue tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) + + Global Instance hashvalue_bool_ne_impl + : HyperBinaryImpl One tint_hashvalue tint_hashvalue tint_bool := { + Hbinary_cond f f' := True; + Hbinary_ocond := otrue2; + Hbinary f f' := negb (hashvalue_eqb f f') + }. + Global Instance hashvalue_bool_ne + : HyperBinaryOp One tint_hashvalue tint_hashvalue tint_bool := { + }. + Proof. + - (* Hbinary_ocond_same *) + reflexivity. + + - (* Hbinary_returns *) + intros; simpl; tauto. + + - (* Hbinary_correct *) + intros f f' v v' fc fc' oc rel rel'. + subst v v'; simpl in *. + constructor; simpl. + unfold hashvalue_eqb. + destruct (hashvalue_eq_dec f f'); + destruct (val_eq_dec (val_of_hashvalue f) (val_of_hashvalue f')); + try reflexivity. + + subst. + congruence. + + contradict n. + auto using val_of_hashvalue_injective. + Qed. + Global Instance int_hashvalue_ne_passthrough + : HyperBinaryPassthrough One tint_hashvalue tint_hashvalue tint_bool. + (* Proof. esplit. + simpl; congruence. + Qed. *) diff --git a/src/core/MemoryModel.v b/src/core/MemoryModel.v new file mode 100755 index 0000000..6250c7f --- /dev/null +++ b/src/core/MemoryModel.v @@ -0,0 +1,139 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** Adapter module aggregating all abstractions in liblayers in CompCertX. + For the adapter _type class_, see [HyperMem]. *) + +(* CompCertX modules *) +(* Require compcertx.common.EventsX. *) + +Require Import ZArith. +Require Import backend.Values.HighValues. +Require Import backend.MemoryModel. +Require Export backend.AbstractData. +(* Require Import liblayers.compat.CompatLayers. *) (* Todo, liblayers-related *) + +(** * Operations on memory states *) +(* This is copied from Compcert/CompcertX as a temporary thing, + but I didn't copy the corresponding axioms. + TODO: think about which operations the generated DeepSEA code actually uses. *) +Class MemoryModelOps (mem: Type) := { + empty: mem; +}. + +Require Export DeepSpec.lib.Lens. + +Section LAYER_SPEC_CLASSES. + +(* All the synthesis functions are in a context containing a LayerSpecClass. + The programmer using DeepSpec will eventually need to provide instances for it, this + contains all the proof obligations for the various theorems. + + Edsger will generate instances in LayerXX.v + + The LayerSpecClass is made up of a double-inheritance diamond + OverLaySpecClass + / \ + / \ + LayerSpecClass HyperMem + \ / + \ / + UnderLaySpecClass / + + There is some complicated technical reason that they need to be split up, which we sadly + no longer remember. + *) + +Class LayerSpecClass : Type := { + memModelOps :> MemoryModelOps mem; + + GetHighData : Type; + GetLowData := GetHighData +}. + +Context`{LayerSpec : LayerSpecClass}. + +Class CompatRelOps (D1 D2: compatdata) := + { + relate_AbData: meminj -> cdata_type D1 -> cdata_type D2 -> Prop; + match_AbData: cdata_type D1 -> mem -> meminj -> Prop; + new_glbl: list AST.ident + }. + +Class CompatRel D1 D2 `{ops: !CompatRelOps D1 D2} := + { + }. + +Class compatrel D1 D2 := + { + crel_ops :> CompatRelOps D1 D2; + crel_prf :> CompatRel D1 D2 + }. + +(* This contains cdata, which is the layer invariant for the abstract data. + + In fact, CompatDataOps contain three things, + the high level invariant, which is specified by the DeepSpec programmer in the .ds source file + "low level invariant" (misnomer) about assembly level stuff. + kernel mode + *) +Class OverlaySpecClass : Type := { + + cdataOpsHigh :> CompatDataOps GetHighData; + cdataHigh :> CompatData GetHighData; + GetHighDataX : compatdata + := @cdata GetHighData _ _ + +}. + +(* This contains the same thing (for the lower layer), and additionally + GetLowLayer, which is the layer interface for the underlay (a map from + identifiers to method semantics). +*) +Class UnderlaySpecClass : Type := { + + cdataOpsLow : CompatDataOps GetLowData; + cdataLow : CompatData GetLowData; + GetLowDataX : compatdata + := cdata (*mem _ *) GetLowData (* _ _ *); + MemLow := mwd GetLowDataX; + +}. + +Context`{OverlaySpec : !OverlaySpecClass}. +Context`{UnderlaySpec : !UnderlaySpecClass}. + +(* The refinement relation (compat_rel gives relate_AbFata and match_Abdata, + and we also define a name mem_match for talking about them together). +*) +Class HyperMem : Type := { + Hcompatrel :> compatrel GetHighDataX GetLowDataX; + mem_match j (d : GetHighData)(ml : MemLow) := + relate_AbData j d (snd ml) /\ match_AbData d (fst ml) j +}. + + +Context`{HM : !HyperMem}. + +End LAYER_SPEC_CLASSES. diff --git a/src/core/SEnv.v b/src/core/SEnv.v new file mode 100755 index 0000000..0da9c07 --- /dev/null +++ b/src/core/SEnv.v @@ -0,0 +1,112 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** Environment and utility definitions for synthesizing *) + +(* CompCert modules *) +Require Import backend.Values.HighValues. +Require Import backend.Environments.AllEnvironments. +Require Import cclib.Maps. +Require Import backend.MemoryModel. +Require Import backend.Ctypes. + +(* DeepSpec modules *) +Require Import DeepSpec.core.Cval. +Require Import DeepSpec.core.HyperType. +Require Import DeepSpec.core.MemoryModel. + +Require Import DeepSpec.lib.SimpleMaps. +Require Export DeepSpec.lib.SimpleIndexedMaps. + +Module TypePairProjection <: TypeProjection. + Definition A := hyper_type_pair. + Definition proj := tp_ft. +End TypePairProjection. + +Module SpecTree := IPList(TypePairProjection). + +Definition spec_env_t := SpecTree.t. + +Definition senv_cond {tmp} (se : spec_env_t tmp) + := forall i tp hti, + forall eq : (tmp ! i)%alist ~~~ Some (@mk_hyper_type_pair tp hti), + ht_ft_cond (SpecTree.get_eq i se eq). + +Definition lenv_cond `{LayerSpecClass}{tmp} (se : spec_env_t tmp)(le : temp_env) + := forall i tp hti, + forall eq : (tmp ! i)%alist ~~~ Some (@mk_hyper_type_pair tp hti), + ht_ft_cond (SpecTree.get_eq i se eq) /\ + ht_rel_some (SpecTree.get_eq i se eq) (le ! i). + +Lemma lenv_senv_cond `{LayerSpecClass}{tmp}{se : spec_env_t tmp}{le} : + lenv_cond se le -> senv_cond se. +Proof. + intros le_cond i tp hti eq. + apply le_cond. +Qed. + +Module HyperTypeProjection <: TypeProjection. + Definition A := hyper_type_pair. + Definition proj htp := HyperType (tp_type_pair htp). +End HyperTypeProjection. + +Definition proj1_ex {A : Prop}{P : A -> Prop}(e : ex P) : A := + match e with ex_intro x _ => x end. +Definition proj2_ex {A : Prop}{P : A -> Prop}(e : ex P) : P (proj1_ex e) := + match e with ex_intro _ Px => Px end. +Definition proj1_sig2 {A P Q}(s : sig2 P Q) : A := + match s with exist2 a _ _ => a end. +Definition proj2_sig2 {A}{P Q : A -> Prop}(s : sig2 P Q) + : P (proj1_sig2 s) := match s with exist2 _ p _ => p end. +Definition proj3_sig2 {A}{P Q : A -> Prop}(s : sig2 P Q) + : Q (proj1_sig2 s) := match s with exist2 _ _ q => q end. + +Definition car := @fst. +Definition cdr := @snd. +Definition cadr A B C a := car B C (cdr A _ a). +Definition cddr A B C a := cdr B C (cdr A _ a). +Definition caar A B C a := car A B (car _ C a). +Definition cdar A B C a := cdr A B (car _ C a). +Definition caddr A B C D a := car C D (cddr A B _ a). +Definition cdddr A B C D a := cdr C D (cddr A B _ a). +Definition cadddr A B C D E a := car D E (cdddr A B C _ a). +Definition cddddr A B C D E a := cdr D E (cdddr A B C _ a). +Definition caddddr A B C D E F a := car E F (cddddr A B C D _ a). +Definition cdddddr A B C D E F a := cdr E F (cddddr A B C D _ a). +Arguments car [A B] _. Arguments cdr [A B] _. +Arguments cadr [A B C] _. Arguments cddr [A B C] _. +Arguments caar [A B C] _. Arguments cdar [A B C] _. +Arguments caddr [A B C D] _. Arguments cdddr [A B C D] _. +Arguments cadddr [A B C D E] _. Arguments cddddr [A B C D E] _. +Arguments caddddr [A B C D E F] _. Arguments cdddddr [A B C D E F] _. + +Definition proj12 A B C a := @proj1 B C (@proj2 A _ a). +Definition proj22 A B C a := @proj2 B C (@proj2 A _ a). +Definition proj122 A B C D a := @proj1 C D (@proj22 A B _ a). +Definition proj222 A B C D a := @proj2 C D (@proj22 A B _ a). +Arguments proj12 [A B C] _. Arguments proj22 [A B C] _. +Arguments proj122 [A B C D] _. Arguments proj222 [A B C D] _. + +Lemma some_injective {A}{a b : A} : Some a = Some b -> a = b. +Proof. injection 1. trivial. Qed. diff --git a/src/core/Syntax.v b/src/core/Syntax.v new file mode 100755 index 0000000..0226bec --- /dev/null +++ b/src/core/Syntax.v @@ -0,0 +1,746 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** Abstract syntax trees for DeepSpec expressions and statements. *) + +(* Standard library modules *) +Require Import BinPosDef. (* positive_scope *) +Require Import List. +Require Import Bool. +Require Import BinInt. + +(* CompCert modules *) +Require Import backend.AST. +(*Require Import compcert.common.Events. *) +Require Import backend.MemoryModel. +Require Import backend.Values.HighValues. +Require Import cclib.Integers. +Require Import cclib.Maps. +Require Import backend.Statements.StmtMiniC. +Require Import backend.Cop. +Require Import backend.Ctypes. +Require Import backend.MachineModel. + +(* DeepSpec modules *) +Require Import DeepSpec.core.Cval. +Require Import DeepSpec.core.SEnv. +Require Import DeepSpec.core.HyperType. +Require Import DeepSpec.core.HyperTypeInst. +Require Import DeepSpec.core.MemoryModel. +Require Import DeepSpec.lib.OProp. +Require Import DeepSpec.lib.Monad.ContOptionMonad. + +Generalizable Variables tp tpl tpr tpo. (* Enabling `{ } and `( ) *) + +(* DeepSpec variables are represented as lenses to the abstract state + a block number. + This is a special case of ltype_pair. *) +Section VARIABLE. + (* Context `{LayerSpec : LayerSpecClass}. *) + Context `{HM : HyperMem}. + Definition variable tp := ltype_pair tp. + + (* TODO: move this to the backend director: *) + Definition offset_zero := Int256.zero. + + Record variable_prf {tp}`{HyperTypeImpl tp}(v : variable tp) : Prop + := mk_var_prf { + (* This is the only difference from the general ltype_pair: we require the offset to be zero. *) + VARoffset_zero : v.(ltype_ident) = Field Global (ident_ext_base (v.(ltype_ident))); + VARltype : HyperLType v + }. + Global Arguments VARoffset_zero {_ _ _} _. + Global Arguments VARltype {_ _ _} _. + + (* This should better be named HyperLTypePassthrough. + For each operation (here, variable reads and writes) we define a Class with the side conditions that are needed to prove that + passthrough methods using that operation are correct. + + Then for each variable in the DeepSpec source program, Edsger will generate an instance. + + These three conditions are better seen as conditions of the refinement relation than on the lens itself. + *) + (* TODO: Sadly, we will probably have to fix this sooner rather than later. *) + Class variable_passthrough_prf {tp}`{HyperTypeImpl tp}(v : variable tp) : Prop + := mk_var_pt_prf { + (* When reading values from related ds, the results are equal modolo injection. *) + VAR_relate_impl_eq j d1 d2 : + relate_AbData j d1 d2 -> + (v.(ltype_get) d1) = (v.(ltype_get) d2); + + (* When writing equal values, the relation still holds. *) + VAR_relate_impl_update j d1 d2 : + relate_AbData j d1 d2 -> + forall f, + relate_AbData j (v.(ltype_set) f d1) (v.(ltype_set) f d2); + + (* Intuitively, this says that writing using the lens does not change the ++++ part of the abstract data. + (Or even if it did change, at least it would still be related.) + + [ raw memory ][++++ abstract data] + [ raw memory :::::][abstract data] + *) + VAR_match_impl_update j d m : + match_AbData d m j -> + forall f, + match_AbData (v.(ltype_set) f d) m j; + (* + (* When writing values to underlay's private variable, overlay's validStateHigh + should be presrved, since it should never metion underlay's private variables. *) + VAR_valid_update d: + validStateHigh d -> + forall f, + validStateHigh (v.(ltype_set)f d) +*) + }. +End VARIABLE. + +Coercion create_type_pair_marker := @create_type_marker type_pair. + +(** * Expressions *) +(** Clight expression is a clone of Clight expression aiming for synthesizing + CertiKOS. It started as a subset of Clight, but diverged in several key + points making it more suitable for synthesizing and proving properties. + - Separated L-value and R-value expressions. + - No addressable local variables; all local variables are temporaries. + - Evaluation of either expression never requires memory access. + - R-value can no longer return structures or unions ([access_mode ty = + By_copy]); arrays and functions are still possible ([access_mode ty = + By_reference]), but this is subject to change in the future. + - No floating point constants as they are not used in kernels. + - Flattened address-of operator ([&]): inlining all potential L-value + expressions. + + The two inductively defined data types, [expr_constr] (expression + construct) and [lexpr_constr] (L-value expression construct), just like + temporaries, are also annotated with the [type_pair] indicating the type + of their evaluation results. They are heavily annotated with type classes + defined in {{DeepSpec.HyperType.html}HyperType}. The type classes play + a huge role in [synth_expr]. + + All branches in [expr_constr] are pretty much the same to their Clight + counterparts, except for [ECaddrof_var], which is the inlined version of + [Eaddrof (Evar i t) t'], and [ECaddrof_pfield], which is the inlined + version of [Eaddrof (Efield (Ederef e t) x t') t'']. +*) +Section EXPR_CONSTR. + Context `{HM : HyperMem}. + (* For instance, lexpr variable definitions contain a lens, so they need to know the type of + the abstract data, which is part of LayerSpecClass, which is implicitly + part of the context because it is a parameter of HM. *) + + (* Annotating (indexing) [ht] instead of parametrizing it to make declaring + subclasses in constructors easier. *) + Inductive expr_constr : forall tp {hti : HyperTypeImpl tp}, Type := + | ECconst_int : forall tp `{HyperTypeImpl tp}, + unpair_ft tp -> Int.int -> expr_constr tp (**r integer literal *) + | ECconst_int256 : forall tp `{HyperTypeImpl tp}, + unpair_ft tp -> Int256.int -> expr_constr tp (**r integer literal *) + | ECtempvar : forall tp `{HyperTypeImpl tp}, + ident -> expr_constr tp (**r temporary variable *) + | ECbuiltin0 : forall tp `{HyperTypeImpl tp, !HyperBuiltin0Impl tp}, + expr_constr tp (**r Ethereum EVM builtin *) + | ECbuiltin1 : forall atp tp `{HyperTypeImpl atp, HyperTypeImpl tp, !HyperBuiltin1Impl atp tp}, + expr_constr atp -> expr_constr tp (**r Ethereum EVM builtin *) + | ECunop : forall tp op `{HyperTypeImpl tp, HyperTypeImpl tpo, HyperUnaryImpl op tpo tp}, + expr_constr tpo -> expr_constr tp + (**r unary operation *) + | ECbinop : forall tp op `{HyperTypeImpl tp, HyperTypeImpl tpl, HyperTypeImpl tpr, + HyperBinaryImpl op tpl tpr tp}, + expr_constr tpl -> expr_constr tpr -> expr_constr tp + (**r binary operation *) +(* Note: this is never generated by Edsger. *) +(* | ECcast : forall tp `{HyperTypeImpl tp, HyperTypeImpl tpo, HyperCastImpl tpo tp}, + expr_constr tpo -> expr_constr tp + (**r type cast ([(ty) e]) *) *) +(* Note: this is never generated by Edsger. *) +(* | ECaddrof : forall {tp}`{HyperTypeImpl tp}, + lexpr_constr tp -> expr_constr tchar_pointer_globalpointer + (* Currently Edsger never generates ECaddrof, but it may be useful in future work. So expr and lexpr are mutually recursive. *) *) +(* Note: this is never generated by Edsger. *) +(* | ECrespec : forall {tp}`{HyperTypeImpl tp} tmp', + expr_constr tint_Z32 -> + forall spec : spec_env_t tmp' -> unpair_ft tp, + expr_constr tp *) + + (** There are only three acceptable L-value expressions and they are + represented by the three constructors of [lexpr_constr]. [LCvar] + and [LCderef] are exactly lixe [Evar] and [Ederef] (note that + [LCderef] takes [expr_constr] as an argument, not [lexpr_constr]: + [lexpr_constr] is not recursive). + + [LCpfield] is equivalent to [Efield (Ederef e t) x t'], which is + inlined in [ECaddrof_pfield]. + *) + with lexpr_constr : forall tp `{hti : HyperTypeImpl tp}, Type := + | LCvar : forall`{HyperTypeImpl tp}(v : variable tp), lexpr_constr tp + (**r variable *) + | LCfield : forall tp id `{HyperTypeImpl tp, HyperTypeImpl tpl, !HyperFieldImpl tpl tp id}, + lexpr_constr tpl -> lexpr_constr tp + (**r access to a member of a struct or union *) + | LCindex : forall tp `{HyperTypeImpl tp, HyperTypeImpl tpl, !HyperIndexImpl tpl tp}, + lexpr_constr tpl -> expr_constr tint_Z32 -> lexpr_constr tp + (**r access to an element of an array. *) + | LChash : forall tp `{HyperTypeImpl tp, HyperTypeImpl tpl, !HyperIntHashImpl tpl tp}, + lexpr_constr tpl -> expr_constr tint_U -> lexpr_constr tp + (**r access to an element of a hashtable. *) + . + + Fixpoint lexpr_is_ghost `{HyperTypeImpl}(e : lexpr_constr tp) := + match e with + | LCvar _ _ v => v.(ltype_ghost) + | LCfield _ tid _ _ _ pfieldi e' => lexpr_is_ghost e' + | LCindex _ _ _ _ _ e_arr e_idx => lexpr_is_ghost e_arr + | LChash _ _ _ _ _ e_arr e_idx => lexpr_is_ghost e_arr + end. + + (* This inductive type contains the verification conditions needed for the corresponding + expression type. Each DeepSpec expression will generate one expr_constr (in LayerFoo.v) + and a tactic will generate an expr_constr_prf (in ObjBarCode.v), where Foo is the layer + containing the object Bar. + + The philosophy is that verification conditions that can be discharged fully automatically + are put into the _prf type, while those that need manual work are put into the _cond . + *) + (* TODO fill in other constructors *) + Inductive expr_constr_prf : forall `{HyperType tp}, expr_constr tp -> Prop := + (* | ECPconst_int : forall`{HyperType tp} f i (fc : ht_ft_cond f) + (rel : ht_cval f = CVval (Vint i)), expr_constr_prf (ECconst_int _ f i) *) + | ECPconst_int256 : forall`{HyperType tp} f i (fc : ht_ft_cond f) + (rel : ht_cval f = CVval (Vint i)), expr_constr_prf (ECconst_int256 _ f i) + | ECPtempvar : forall`{ht : HyperType tp} t, expr_constr_prf (ECtempvar tp t) + + | ECPbuiltin0 : forall `{HyperBuiltin0 (LayerSpec := LayerSpec) tp, !HyperType tp}, expr_constr_prf (ECbuiltin0 tp) + | ECPbuiltin1 : forall atp `{HyperBuiltin1 (LayerSpec := LayerSpec) atp tp, !HyperType atp, !HyperType tp} e, + expr_constr_prf e -> + expr_constr_prf (ECbuiltin1 atp tp e) + + | ECPunop : forall op `{HyperUnaryOp op tpo tp, ht : !HyperType tp, + hto : !HyperType tpo} e, + expr_constr_prf e -> expr_constr_prf (ECunop _ op e) + | ECPbinop : forall op `{HyperBinaryOp op tpl tpr tp, ht : !HyperType tp, + htl : !HyperType tpl, htr : !HyperType tpr} el er, + expr_constr_prf el -> expr_constr_prf er -> + expr_constr_prf (ECbinop _ op el er) + + with lexpr_constr_prf : forall`{HyperType tp}, lexpr_constr tp -> Prop := + | LCPvar : forall`{HyperType tp} var (v_prf : variable_prf var), + lexpr_constr_prf (LCvar var) + | LCPfield : forall id `{HyperField tpl tp id, !HyperType tpl, !HyperType tp} + e (ec : lexpr_constr_prf e), + lexpr_constr_prf (LCfield tp id e) + | LCPindex : forall`{HyperIndex tpl tp, !HyperType tp, !HyperType tpl} e idx, + lexpr_constr_prf e -> expr_constr_prf idx -> + lexpr_constr_prf (LCindex tp e idx) + | LCPhash : forall`{HyperIntHash tpl tp, !HyperType tp, !HyperType tpl} e idx, + lexpr_constr_prf e -> expr_constr_prf idx -> + lexpr_constr_prf (LChash tp e idx) + . + + (* Each DeepSpec thing also defines a corresponding _passthrough_prf. + + The _prf is used for proving correctness when a primitive is first added to a layer, and + the _passthrough_prf is used when proving correctness of higher layers when the primitive is + passed through. + *) + + Inductive expr_constr_passthrough_prf : + forall`{HyperTypeImpl tp}, expr_constr tp -> Prop := + | ECPPconst_int : forall`{hti : HyperTypeImpl tp} f i, + expr_constr_passthrough_prf (ECconst_int tp f i) + | ECPPtempvar : forall`{hti : HyperTypeImpl tp} t, + expr_constr_passthrough_prf (ECtempvar tp t) + | ECPPunop : forall tp op `{htu : HyperUnaryPassthrough op tpo tp} e, + expr_constr_passthrough_prf e -> expr_constr_passthrough_prf (ECunop tp op e) + | ECPPbinop : forall tp op `{htb : HyperBinaryPassthrough op tpl tpr tp} el er, + expr_constr_passthrough_prf el -> expr_constr_passthrough_prf er -> + expr_constr_passthrough_prf (ECbinop tp op el er) + + with lexpr_constr_passthrough_prf : + forall`{HyperTypeImpl tp}, lexpr_constr tp -> Prop := + | LCPPvar : forall`{hti : HyperTypeImpl tp} var (v_prf : variable_passthrough_prf var), + lexpr_constr_passthrough_prf (LCvar var) + | LCPPfield : forall id `{htf : HyperFieldPassthrough tpl tp id} e, + lexpr_constr_passthrough_prf e -> + lexpr_constr_passthrough_prf (LCfield tp id e) + | LCPPindex : forall `{htd : HyperIndexPassthrough tpl tp} e idx, + lexpr_constr_passthrough_prf e -> expr_constr_passthrough_prf idx -> + lexpr_constr_passthrough_prf (LCindex tp e idx) + . +End EXPR_CONSTR. + +Section STMT_CONSTR. + (** * Commands *) + (** Again, Clight statement is a clone of Clight statement but wildly + modified. Key differences include: + - Each statement can either read, write, or do not access the memory, + but never read _and_ write. + - Memory accessing has been limited to one primitive data type at a time. + (No whole structure or union assignments.) + - The statements are indexed by wheather they will [return] or not. + This is a must analysis, not a may analysis. + - Single assignments on all temporaries. (Future versions may have + special forms of assignments for loops and branching.) + + In the fine-grain abstract layer model, function calls are restricted + to calls to lower layers marked as external function calls. Intra-layer + function calls indicate the need to split a layer into multiple ones, + and recursion is (semi-)forbidden for stack size concern. + *) + (* Context `{LayerSpec : LayerSpecClass}. *) + Context `{HM : HyperMem}. + + Definition function_return_dec returns + := match tp_ty returns with + | Tvoid => false + | _ => true + end. + + Require Import DeepSpec.lib.Monad.StateMonadOption. + Import MonadNotation. + Open Scope monad_scope. + Definition DS := osT GetHighData. + Instance Monad_DS : Monad DS := MosT GetHighData. + + Instance MonadLaws_DS : MonadLaws (Monad_stateT GetHighData Monad_option). + Proof. + apply MonadLaws_MosT. + Qed. + + Instance MonadState_DS : MonadState GetHighData DS. + apply MonadState_stateT. + apply Monad_option. + Defined. + + Instance MonadZero_DS : MonadZero DS. + apply MonadZero_stateT. + apply Monad_option. + apply Zero_option. + Defined. + + + (* The underlay interface for a primitive is a Coq Prop, a relation from memory etc to the final value. + But we want a Coq function. So the primitive record contains the function. + + Then the class primitive_exec_prf says that the primitive is equivalent to the specification in the + underlay interface. + *) + Record primitive (argt : list hyper_type_pair)(ret : hyper_type_pair) + : Type := mk_prim { + PRIMident : ident; + (* Todo: maybe we need this for the C backend. If so, need to parameterize it as another opaque type in MiniC.v *) + (*PRIMcc : calling_convention; (* Either C or Assembly calling convention. This is not the same as the [[semantic]] in the source file! *)*) + PRIMghost : bool; (* These track the source file. *) + PRIMpure : bool; (* This is called "const" in the source file. *) + + PRIMargt_marker : type_marker argt; + PRIMret_marker : type_marker ret; + + + (* PRIMsem_opt, the "monadic" version, combines PRIMcond and PRIMsem into one thing. + There is a verification condition below saying that they must be equivalent; + the synthesis use them interchangably for convenience. *) + PRIMcond : HList tp_ft argt -> machine_env GetHighData -> GetHighData -> Prop; + + PRIMsem_opt : HList tp_ft argt -> machine_env GetHighData -> DS (tp_ft ret) +(* PRIMsem : HList tp_ft argt -> GetHighData -> + GetHighData * (* trace * *) tp_ft ret; + PRIMsem_opt : HList tp_ft argt -> GetHighData -> + ContOpt (GetHighData * (* trace * *) tp_ft ret) *) + }. + (* begin hide *) + Global Arguments PRIMident {_ _} _. + (* Global Arguments PRIMcc {_ _} _. *) + Global Arguments PRIMghost {_ _} _. + Global Arguments PRIMpure {_ _} _. + Global Arguments PRIMcond {_ _} _ _ _. + (* Global Arguments PRIMsem {_ _} _ _ _. *) + (* Global Arguments PRIMsem_opt {_ _} _ _ _ _ _. *) + Global Arguments PRIMsem_opt {_ _} _ _. + (* end hide *) + + Fixpoint to_typelist argt := + match argt with + | nil => Ctypes.Tnil + | cons x xs => Ctypes.Tcons (tp_ty x) (to_typelist xs) + end. + + (* This class has verification conditions for all primitives (ghost and non-ghost). *) + Class primitive_prf {argt returns}(prim : primitive argt returns) := mk_prim_prf { + PRIMret_cond : + forall args me, + ht_list_ft_cond args -> + ht_list_valid_ft_cond args -> + forall (s : GetLowData) v, + evalStateT (prim.(PRIMsem_opt) args me) s = ret v -> + high_level_invariant (CompatDataOps := cdataOpsLow) s -> + (* validStateLow d -> *) + (* (exists j m, match_AbData d m j) -> *) + ht_ft_cond v /\ + (function_return_dec returns = true -> ht_valid_ft_cond v); + PRIMis_pure : + prim.(PRIMpure) = true -> + forall args me d, + forall d', execStateT (prim.(PRIMsem_opt) args me) d = ret d' -> + d = d'; + + PRIMinv : + prim.(PRIMpure) = false -> + forall args me, + ht_list_ft_cond args -> + ht_list_valid_ft_cond args -> + forall (s s' : GetLowData), + execStateT (prim.(PRIMsem_opt) args me) s = ret s' -> + high_level_invariant (CompatDataOps := cdataOpsLow) s -> + high_level_invariant (CompatDataOps := cdataOpsLow) s' + }. + + (* This class contains verification conditions specific to non-ghost primitives. *) + Class primitive_exec_prf {argt ret}(prim : primitive argt ret) := mk_prim_exec_prf { + }. + + Class primitive_passthrough_prf {argt ret}(prim : primitive argt ret) + := mk_prim_pt_prf { + }. + + Definition expr_constr_list := + HList (fun htp => expr_constr (tp_type_pair htp)). + Definition expr_constr_prf_conj {ls}(es : expr_constr_list ls) := + HList_fold_right_nodep (fun htp e p => + p /\ { ht : HyperType (tp_type_pair htp) | expr_constr_prf (H := ht) e }) + True es. + + Definition expr_constr_passthrough_prf_conj {ls}(es : expr_constr_list ls) := + HList_fold_right_nodep (fun htp e p => p /\ expr_constr_passthrough_prf e) + True es. + + Inductive cmd_constr : forall (returns : hyper_type_pair), Type := + | CCskip : cmd_constr void_unit_pair (**r do nothing *) + | CClet : forall {r}`{HyperTypeImpl tp}(id : ident), + cmd_constr (mk_hyper_type_pair tp) -> cmd_constr r -> cmd_constr r + (**r binding [let x = s1 in s2] *) + | CCload : forall`{HyperTypeImpl tp}, + lexpr_constr tp -> cmd_constr (mk_hyper_type_pair tp) + (**r memory read *) + | CCstore : forall`{HyperTypeImpl tp}, + lexpr_constr tp -> expr_constr tp -> cmd_constr void_unit_pair + (**r memory write [lvalue = rvalue] *) + | CCsequence : forall {r}, + cmd_constr void_unit_pair -> cmd_constr r -> cmd_constr r (**r sequence *) + | CCifthenelse : forall {r}, + expr_constr tint_bool -> cmd_constr r -> cmd_constr r -> cmd_constr r + (**r conditional *) + | CCfor : forall id_it id_end : ident, + expr_constr tint_Z32 -> expr_constr tint_Z32 -> cmd_constr void_unit_pair -> + cmd_constr void_unit_pair (**r bounded loop *) + | CCfirst : forall {r}(id_it id_end id_dest : ident), + expr_constr tint_Z32 -> expr_constr tint_Z32 -> + cmd_constr int_bool_pair -> + cmd_constr r -> cmd_constr r -> cmd_constr r (**r linear search *) + | CCfold : forall `{HyperTypeImpl tp}(id_it id_end id_recur id_dest : ident), + expr_constr tint_Z32 -> expr_constr tint_Z32 -> expr_constr tp -> + cmd_constr (mk_hyper_type_pair tp) -> + cmd_constr (mk_hyper_type_pair tp) (**r linear scan *) + | CCcall : forall {argt ret}, + primitive argt ret -> expr_constr_list argt -> cmd_constr ret + (**r (internal object) function call *) + (* | CCcall_ext : forall {argt ret}, + int256 -> (* the address of the external object/contract to call *) + primitive argt ret -> expr_constr_list argt -> cmd_constr ret + (**r (external contract) function call *) *) + | CCyield : forall`{HyperTypeImpl tp}, + expr_constr tp -> cmd_constr (mk_hyper_type_pair tp) + (**r yield or return a value depends on context *) + + (**r Belows are DeepSpec specific *) + | CCconstr : forall`{HyperTypeImpl tp}(fld_ids : list ident) fld_tps, + lexpr_constr tp -> expr_constr_list fld_tps -> + list_curried_type fld_tps (unpair_ft tp) -> + cmd_constr void_unit_pair + | CCassert : cmd_constr int_bool_pair -> cmd_constr void_unit_pair + (**r specification only assertion *) + | CCdeny : cmd_constr int_bool_pair -> cmd_constr void_unit_pair + (**r specification only assertion *) + (* | CCghost : cmd_constr void_unit_pair -> cmd_constr void_unit_pair *) + (**r specification only command *) + | CCpanic : forall tp`{HyperTypeImpl tp}, + cmd_constr (mk_hyper_type_pair tp) + (**r specification only failure (w/ a dummy value) *) + (* There are two versions of of respec, which take a monadic or a non-monadic functional specification. + The difference is that CCrespec only can encode pure specifications (it doesn't return a new memory). + So if you want a non-pure specification you have to use CCrespec_opt, + and if you want to use a respec in a loop test you have to use CCrespec. + *) + | CCrespec : forall {r} tmp', cmd_constr r -> + forall spec : spec_env_t tmp' -> DS (tp_ft r), + cmd_constr r (**r manually specify the behavior *) + | CCrespec_opt : forall {r} tmp', cmd_constr r -> + forall spec : machine_env GetHighData -> spec_env_t tmp' -> DS (tp_ft r), (* GetHighData -> *) + (* ContOpt (GetHighData * (* trace * *) tp_ft r), *) + cmd_constr r (**r manually specify the behavior *) + . + + Inductive cmd_constr_passthrough_prf : forall ret, cmd_constr ret -> Prop := + | CCPPskip : cmd_constr_passthrough_prf _ CCskip + | CCPPlet : forall {r}`{hti : HyperTypeImpl tp}(id : ident) c1 c2, + cmd_constr_passthrough_prf (mk_hyper_type_pair tp) c1 -> + cmd_constr_passthrough_prf r c2 -> + cmd_constr_passthrough_prf r (CClet id c1 c2) + | CCPPload : forall`{hti : HyperTypeImpl tp} e, + lexpr_constr_passthrough_prf e -> + cmd_constr_passthrough_prf (mk_hyper_type_pair tp) (CCload e) + | CCPPstore : forall`{hti : HyperTypeImpl tp} el er, + lexpr_constr_passthrough_prf el -> expr_constr_passthrough_prf er -> + cmd_constr_passthrough_prf _ (CCstore el er) + | CCPPsequence : forall {r} c1 c2, + cmd_constr_passthrough_prf _ c1 -> cmd_constr_passthrough_prf r c2 -> + cmd_constr_passthrough_prf r (CCsequence c1 c2) + | CCPPifthenelse : forall {r} e c1 c2, + expr_constr_passthrough_prf e -> + cmd_constr_passthrough_prf r c1 -> cmd_constr_passthrough_prf r c2 -> + cmd_constr_passthrough_prf r (CCifthenelse e c1 c2) + | CCPPfor : forall id_it id_end e1 e2 c, + expr_constr_passthrough_prf e1 -> expr_constr_passthrough_prf e2 -> + cmd_constr_passthrough_prf _ c -> + cmd_constr_passthrough_prf _ (CCfor id_it id_end e1 e2 c) + | CCPPfirst : forall {r} id_it id_end id_dest e1 e2 c3 c4 c5, + expr_constr_passthrough_prf e1 -> expr_constr_passthrough_prf e2 -> + cmd_constr_passthrough_prf _ c3 -> + cmd_constr_passthrough_prf r c4 -> cmd_constr_passthrough_prf r c5 -> + cmd_constr_passthrough_prf r (CCfirst id_it id_end id_dest e1 e2 c3 c4 c5) + | CCPPfold : forall `{hti : HyperTypeImpl tp} id_it id_end id_recur id_dest e1 e2 e3 c, + expr_constr_passthrough_prf e1 -> expr_constr_passthrough_prf e2 -> + expr_constr_passthrough_prf e3 -> + cmd_constr_passthrough_prf (mk_hyper_type_pair tp) c -> + cmd_constr_passthrough_prf (mk_hyper_type_pair tp) + (CCfold id_it id_end id_recur id_dest e1 e2 e3 c) + | CCPPcall : forall {argt ret} prim arg, + expr_constr_passthrough_prf_conj arg -> + primitive_passthrough_prf prim -> + cmd_constr_passthrough_prf _ (@CCcall argt ret prim arg) + (* external call *) + | CCPPcall_ext : forall {argt ret} prim arg, + expr_constr_passthrough_prf_conj arg -> + primitive_passthrough_prf prim -> + cmd_constr_passthrough_prf _ (@CCcall argt ret prim arg) + | CCPPyield : forall`{hti : HyperTypeImpl tp} e, + expr_constr_passthrough_prf e -> + cmd_constr_passthrough_prf (mk_hyper_type_pair tp) (CCyield e) + | CCPPconstr : forall`{HyperTypeImpl tp} fld_ids fld_tps el flds constr + `{hcp : !HyperConstructorPassthrough tp fld_tps constr}, + lexpr_constr_passthrough_prf el -> expr_constr_passthrough_prf_conj flds -> + cmd_constr_passthrough_prf _ (CCconstr fld_ids fld_tps el flds constr) + | CCPPassert : forall c, + cmd_constr_passthrough_prf _ c -> cmd_constr_passthrough_prf _ (CCassert c) + | CCPPdeny : forall c, + cmd_constr_passthrough_prf _ c -> cmd_constr_passthrough_prf _ (CCdeny c) +(* | CCPPghost : forall c, + cmd_constr_passthrough_prf _ c -> cmd_constr_passthrough_prf _ (CCghost c) *) + | CCPPpanic : forall tp`{HyperTypeImpl tp}, + cmd_constr_passthrough_prf (mk_hyper_type_pair tp) (CCpanic tp) + | CCPPrespec : forall {r} tmp' c spec, + cmd_constr_passthrough_prf r (CCrespec tmp' c spec) + | CCPPrespec_opt: forall {r} tmp' c spec, + cmd_constr_passthrough_prf r (CCrespec_opt tmp' c spec) + . + + (* Local identifier allocation: + + [FC_ident_start]: first valid local ident -- reserved for return value; + [FC_param_ident_start]: ident for the first function parameter (or + the first local variable if there is none). + + All [SClet]-bindings in [FC_body] should have idents being at least + [FC_param_ident_start + length FC_params]. + *) + Record function_constr : Type := mk_function_constr { + FC_ident_start : ident; + FC_param_ident_start := Pos.succ FC_ident_start; + + FC_params : list hyper_type_pair; + FC_returns : hyper_type_pair; + + FC_body : cmd_constr FC_returns + }. + + Definition function_constr_passthrough_prf f : Prop := + cmd_constr_passthrough_prf f.(FC_returns) f.(FC_body). + +End STMT_CONSTR. + +(** For primitive type markers *) +Coercion create_list_marker := fun A => @create_type_marker (list A). +Coercion create_hyper_type_pair_marker := @create_type_marker hyper_type_pair. + +Section CONSTR_PRF. + Context `{HM : HyperMem}. + + (* The cmd_constr_prf is parametrized by the is_realizing bit. + If this is true, the _prf condition is stronger, since it also includes + conditions needed to prove the synthesized code correct. *) + Inductive cmd_constr_prf (is_realizing : bool) : + forall ret (htr : HyperType (tp_type_pair ret)), cmd_constr ret -> Prop := + | CCPskip : cmd_constr_prf is_realizing _ _ CCskip + | CCPlet : forall r `{ht : HyperType tp} htr id c1 c2, + function_return_dec (mk_hyper_type_pair tp) = true -> + cmd_constr_prf is_realizing (mk_hyper_type_pair tp) ht c1 -> + cmd_constr_prf is_realizing r htr c2 -> + cmd_constr_prf is_realizing r htr (CClet id c1 c2) + | CCPload : forall `{ht : HyperType tp, hbv : !HyperByValueType tp} e + (* {hltrc: HyperLTypeRetCond} *) + (ec : lexpr_constr_prf e), + (* To be realizable, must not read from a ghost value: + is_realizing => ~ is_ghost + <=> + ~ is_realizing \/ ~ is_ghost + <=> + ~ (is_realizing /\ is_ghost) + *) + (is_realizing && lexpr_is_ghost e = false) -> + cmd_constr_prf is_realizing (mk_hyper_type_pair tp) ht (CCload e) + | CCPstore : forall `{ht : HyperType tp, hbv : !HyperByValueType tp} el er + (ecl : lexpr_constr_prf el)(ecr : expr_constr_prf er), + (* To skip realization, must act ghostly not touching real data: + ~ is_realizing => is_ghost + <=> + ~ ~ is_realizing \/ is_ghost + *) + (is_realizing || lexpr_is_ghost el = true) -> + cmd_constr_prf is_realizing _ _ (CCstore el er) + | CCPsequence : forall r htr c1 c2, + cmd_constr_prf is_realizing _ _ c1 -> + cmd_constr_prf is_realizing r htr c2 -> + cmd_constr_prf is_realizing r htr (CCsequence c1 c2) + | CCPifthenelse : forall r htr e c_true c_false + (ec : expr_constr_prf e), + cmd_constr_prf is_realizing r htr c_true -> + cmd_constr_prf is_realizing r htr c_false -> + cmd_constr_prf is_realizing r htr (CCifthenelse e c_true c_false) + | CCPfor : forall id_it id_end e1 e2 c + (ec1 : expr_constr_prf e1)(ec2 : expr_constr_prf e2), + cmd_constr_prf is_realizing _ _ c -> + cmd_constr_prf is_realizing _ _ (CCfor id_it id_end e1 e2 c) + | CCPfirst : forall r htr id_it id_end id_dest e1 e2 c3 c4 c5, + expr_constr_prf e1 -> expr_constr_prf e2 -> + cmd_constr_prf is_realizing int_bool_pair int_bool c3 -> + cmd_constr_prf is_realizing r htr c4 -> + cmd_constr_prf is_realizing r htr c5 -> + cmd_constr_prf is_realizing r htr + (CCfirst id_it id_end id_dest e1 e2 c3 c4 c5) + | CCPfold : forall `{ht : HyperType tp} id_it id_end id_recur id_dest + e1 e2 e3 c (ec1 : expr_constr_prf e1)(ec2 : expr_constr_prf e2) + (ec3 : expr_constr_prf e3), + function_return_dec (mk_hyper_type_pair tp) = true -> + cmd_constr_prf is_realizing (mk_hyper_type_pair tp) ht c -> + cmd_constr_prf is_realizing (mk_hyper_type_pair tp) ht + (CCfold id_it id_end id_recur id_dest e1 e2 e3 c) + | CCPcall : forall argt ret htr prim arg + (IHprim : @primitive_prf _ _ argt ret prim) + (IHprim_exec : primitive_exec_prf prim) + (IHprim_b : primitive_passthrough_prf prim) + (ecs : expr_constr_prf_conj arg), + (* To be realizable, must not receive value from a ghost primitive call: + is_realizing /\ ghost => ~ return + <=> + ~ (is_realizing /\ ghost) \/ ~ return + <=> + ~ (is_realizing /\ return /\ ghost) + + To skip realization, must act ghostly and only call ghost or pure + primitives: + ~ is_realizing => ghost \/ pure + <=> + ~ ~ is_realizing \/ (ghost \/ pure) + *) + (* (is_realizing && function_return_dec ret && prim.(PRIMghost) = false) -> *) + (* (is_realizing || prim.(PRIMghost) || prim.(PRIMpure) = true) -> *) + is_realizing || prim.(PRIMpure) = true -> + cmd_constr_prf is_realizing ret htr (CCcall prim arg) + (* | CCPcall_ext : forall argt ret htr addr prim arg + (IHprim : @primitive_prf _ _ argt ret prim) + (IHprim_exec : primitive_exec_prf prim) + (IHprim_b : primitive_passthrough_prf prim) + (ecs : expr_constr_prf_conj arg), + (* To be realizable, must not receive value from a ghost primitive call: + is_realizing /\ ghost => ~ return + <=> + ~ (is_realizing /\ ghost) \/ ~ return + <=> + ~ (is_realizing /\ return /\ ghost) + + To skip realization, must act ghostly and only call ghost or pure + primitives: + ~ is_realizing => ghost \/ pure + <=> + ~ ~ is_realizing \/ (ghost \/ pure) + *) + (* (is_realizing && function_return_dec ret && prim.(PRIMghost) = false) -> *) + (* (is_realizing || prim.(PRIMghost) || prim.(PRIMpure) = true) -> *) + is_realizing || prim.(PRIMpure) = true -> + cmd_constr_prf is_realizing ret htr (CCcall_ext addr prim arg) *) + | CCPyield : forall `{ht : HyperType tp} e (ec : expr_constr_prf e), + cmd_constr_prf is_realizing (mk_hyper_type_pair tp) ht (CCyield e) + | CCPconstr : forall `{ht : HyperType tp} fld_ids fld_tps el flds constr + `{hc : !HyperConstructor tp fld_ids fld_tps constr} + (ecl : lexpr_constr_prf el)(fldc : expr_constr_prf_conj flds) + (flds_byvalue : HList + (fun htp : hyper_type_pair => + HyperByValueType (tp_type_pair htp)) fld_tps), + (* See [CCPstore] *) + (is_realizing || lexpr_is_ghost el = true) -> + (length fld_ids <= Z.to_nat Int256.modulus)%nat -> + cmd_constr_prf is_realizing _ _ (CCconstr fld_ids fld_tps el flds constr) + + (* These three constructors is where the is_realizing bit changes, their subexpressions are ghost. *) + | CCPassert : forall c, cmd_constr_prf is_realizing _ _ c -> cmd_constr_prf is_realizing _ _ (CCassert c) + | CCPdeny : forall c, cmd_constr_prf is_realizing _ _ c -> cmd_constr_prf is_realizing _ _ (CCdeny c) +(* | CCPghost : forall c, cmd_constr_prf false _ _ c -> cmd_constr_prf is_realizing _ _ (CCghost c) *) + | CCPpanic : forall `{ht : HyperType} (*f*), cmd_constr_prf is_realizing _ _ (CCpanic tp (*f*)) + | CCPrespec : forall r htr tmp' c spec, + (* Not sure if necessary + function_return_dec r = true -> + *) + cmd_constr_prf is_realizing r htr c -> + cmd_constr_prf is_realizing r htr (CCrespec tmp' c spec) + | CCPrespec_opt : forall r htr tmp' c spec, + (* Do not know if [spec] acts on the abstract state stealthily, restrict + it to be realizing. *) + is_realizing = true -> + cmd_constr_prf is_realizing r htr c -> + cmd_constr_prf is_realizing r htr (CCrespec_opt tmp' c spec) + . + + Record function_constr_prf (is_realizing : bool) (f : function_constr) : Prop := mk_function_constr_prf { + (* FC_returns_cast : if function_return_dec f.(FC_returns) + then HyperArgRet (tp_type_pair f.(FC_returns)) + else True; *) + + FC_returns_ht : HyperType (tp_type_pair f.(FC_returns)); + FC_body_prf : cmd_constr_prf is_realizing f.(FC_returns) FC_returns_ht f.(FC_body) + }. + (* begin hide *) + (* Global Arguments FC_returns_cast _ {_} _. + Global Arguments FC_returns_ht _ {_} _. + Global Arguments FC_body_prf _ {_} _. *) + (* end hide *) +End CONSTR_PRF. diff --git a/src/core/SynthesisExpr.v b/src/core/SynthesisExpr.v new file mode 100755 index 0000000..e74ef6c --- /dev/null +++ b/src/core/SynthesisExpr.v @@ -0,0 +1,498 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** Synthesizing Clight expressions, functional expressions, and + specifications from DeepSpec along with the proofs. *) + +(* Standard library modules *) +Require Import BinInt. + +(* CompCert modules *) +Require Import backend.AST. +(*Require Import compcert.common.Events. +Require Import compcert.common.Globalenvs. (* Genv *) *) +Require Import backend.Values.HighValues. +Require Import cclib.Integers. +Require Import cclib.Maps. (* PTree *) +Require Import backend.Expressions.ExpMiniC. +Require Import backend.Expressions.SemanticsMiniC. +Require Import backend.Statements.StmtMiniC. +Require Import backend.Cop. +Require Import backend.Ctypes. +Require Import backend.MachineModel. + +(* DeepSpec modules *) +Require Import DeepSpec.lib.SimpleMaps. +Require Import DeepSpec.lib.SimpleIndexedMaps. +Require Import DeepSpec.lib.OProp. +Require Import DeepSpec.lib.Monad.ContOptionMonad. +Require Import DeepSpec.core.Cval. +Require Import DeepSpec.core.SEnv. +Require Import DeepSpec.core.HyperType. +Require Import DeepSpec.core.HyperTypeInst. +(*Require Import DeepSpec.core.HyperMem.*) +Require Import DeepSpec.core.MemoryModel. +Require Import DeepSpec.core.Syntax. + +Require Import DeepSpec.backend.phase.MiniC.Semantics. + +(*Generalizable Variables tp tpo. (* Enabling `{ } and `( ) *)*) + +(* Prevent use_convert_proper from making setoid rewriting very slow. *) +(* Ltac liblayers.lib.LogicalRelations.use_convert_proper ::= fail. *) (* Todo, liblayers-related *) + +Section EXPR_FUNC. + (* Context `{LayerSpec : LayerSpecClass}. *) + Context`{HM : HyperMem}. + Context {ctx: int256}. + Context (me : machine_env GetHighData). + + Let genv_t := genv. + (* Let env := PTree.empty _ . *) + + (* This is like an extra typecheck. + If we tried to encode the entire type system for the typed intermediate langauge into the expr_constr + datatype, Coq gets very slow. + + The main thing here is that we check that tempvars are well typed with respect to the typing environment tmp. *) + Fixpoint synth_expr_wellformed + (tmp : AList.t hyper_type_pair) + `{HyperTypeImpl} + (e : expr_constr tp) {struct e} : Type := + match e with + | ECconst_int _ _ f i => True + | ECconst_int256 _ _ f i => True + | ECtempvar _ _ i => (tmp ! i ~~~ Some (mk_hyper_type_pair tp))%alist + | ECbuiltin0 _ _ _ => True + | ECbuiltin1 _ _ _ _ _ e' => synth_expr_wellformed tmp e' + | ECunop _ op _ tpo _ _ e' => synth_expr_wellformed tmp e' + | ECbinop _ op _ tpl tpr _ _ _ el er => + synth_expr_wellformed tmp el * synth_expr_wellformed tmp er + end%type + + with synth_lexpr_wellformed tmp `{HyperTypeImpl}(e : lexpr_constr tp) {struct e}: Type := + match e with + | LCvar _ _ v => True + | LCfield _ x _ _ _ pfieldi e' => synth_lexpr_wellformed tmp e' + | LCindex _ _ _ _ _ e' idx => + synth_lexpr_wellformed tmp e' * synth_expr_wellformed tmp idx + | LChash _ _ _ _ _ e' idx => + synth_lexpr_wellformed tmp e' * synth_expr_wellformed tmp idx + end%type. + + (* the synth_expr_ functions generate C expressions. *) + Fixpoint synth_expr_expr + (tmp : AList.t hyper_type_pair) + `{HyperTypeImpl} + (e : expr_constr tp) : expr := + match e with + | ECconst_int _ _ _ i => Econst_int i (unpair_ty tp) + | ECconst_int256 _ _ _ i => Econst_int256 i (unpair_ty tp) + | ECtempvar _ _ tid => Etempvar tid (unpair_ty tp) + | ECbuiltin0 _ _ _ => Ecall0 Hbuiltin0 (unpair_ty tp) + | ECbuiltin1 _ _ _ _ _ e' => Ecall1 Hbuiltin1 (synth_expr_expr tmp e') (unpair_ty tp) + | ECunop _ op _ _ _ _ e' => + Eunop op (synth_expr_expr tmp e') (unpair_ty tp) + | ECbinop _ op _ _ _ _ _ _ el er => + Ebinop op (synth_expr_expr tmp el) (synth_expr_expr tmp er) (unpair_ty tp) + end. + + Fixpoint synth_lexpr_expr tmp `{HyperTypeImpl}(e : lexpr_constr tp) : expr := + match e with + | LCvar _ _ v => Eglob (ident_ext_base v.(ltype_ident)) (unpair_ty tp) + | LCfield _ x _ _ _ pfieldi e' => + Efield (synth_lexpr_expr tmp e') x (unpair_ty tp) + | LCindex _ _ _ _ _ e' idx => + Eindex (synth_lexpr_expr tmp e') (synth_expr_expr tmp idx) (unpair_ty tp) + | LChash _ _ _ _ _ e' idx => + Eindex (synth_lexpr_expr tmp e') (synth_expr_expr tmp idx) (unpair_ty tp) + (* The hash becomes an index into a global array, though we should actually + treat it like that in synthesis, too. *) + end. + + (* the synth_ _spec are the desugaring into a functional Gallina spec. + rexpressions just return a value in the "functional" part of the type pair tp. *) + Fixpoint synth_expr_spec + (tmp : AList.t hyper_type_pair) + `{HyperTypeImpl} + (e : expr_constr tp) + : synth_expr_wellformed tmp e -> spec_env_t tmp -> unpair_ft tp := + match e with + | ECconst_int _ _ f _ => fun _ _ => f + | ECconst_int256 _ _ f _ => fun _ _ => f + | ECtempvar _ _ i => fun wf se => SpecTree.get_eq i se wf + | ECbuiltin0 _ _ _ => fun wf se => Hquery0 me + | ECbuiltin1 _ _ _ _ _ e' => fun wf se => Hquery1 me (synth_expr_spec tmp e' wf se) + | ECunop _ op _ tpo _ _ e' => + fun wf se => Hunary (synth_expr_spec tmp e' wf se) + | ECbinop _ op _ tpl tpr _ _ _ el er => + fun wf se => Hbinary (synth_expr_spec tmp el (car wf) se) + (synth_expr_spec tmp er (cdr wf) se) + end. + + (* lexpressions are desugared into lenses. *) + Fixpoint synth_lexpr_spec tmp `{HyperTypeImpl} (e : lexpr_constr tp) + : synth_lexpr_wellformed tmp e -> spec_env_t tmp -> ltype_pair tp := + match e with + | LCvar _ _ v => fun _ _ => v + | LCfield _ tid _ _ _ pfieldi e' => fun wf se => + let ptr := synth_lexpr_spec tmp e' wf se in + field_ltype_pair ptr tid + | LCindex _ _ _ _ _ e_arr e_idx => fun wf se => + let ptr := synth_lexpr_spec tmp e_arr (car wf) se in + indexing_ltype_pair ptr (synth_expr_spec tmp e_idx (cdr wf) se) + | LChash _ _ _ _ _ e_arr e_idx => fun wf se => + let ptr := synth_lexpr_spec tmp e_arr (car wf) se in + inthash_ltype_pair ptr (synth_expr_spec tmp e_idx (cdr wf) se) + end. + + (* TODO: the expression synthesis is probably the only place that uses the OProp machinery, + but actually it is probably not needed, and we could simplify the code by just using Props. *) + + (* This is the "synthesis safety condition". It will eventually be part of the verification condition + to be proven in FooProof.v *) + Fixpoint synth_expr_ocond + (tmp : AList.t hyper_type_pair) + `{HyperTypeImpl} + (e : expr_constr tp) + : synth_expr_wellformed tmp e -> OProp1 (spec_env_t tmp) := + match e with + | ECconst_int _ _ f i => fun _ => otrue1 + | ECconst_int256 _ _ f i => fun _ => otrue1 + | ECtempvar _ _ t => fun _ => otrue1 + | ECbuiltin0 _ _ _ => fun _ => otrue1 + | ECbuiltin1 _ _ _ _ _ e' => fun wf => synth_expr_ocond tmp e' wf + | ECunop _ op _ _ _ _ e' => fun wf => + synth_expr_ocond tmp e' wf /\ + Hunary_ocond m{ synth_expr_spec tmp e' wf } + | ECbinop _ op _ _ _ _ _ _ el er => fun wf => + synth_expr_ocond tmp el (car wf) /\ + synth_expr_ocond tmp er (cdr wf) /\ + Hbinary_ocond m{ synth_expr_spec tmp el (car wf) } + m{ synth_expr_spec tmp er (cdr wf) } + end%oprop1 + + with synth_lexpr_ocond tmp `{HyperTypeImpl}(e : lexpr_constr tp) + : synth_lexpr_wellformed tmp e -> OProp1 (spec_env_t tmp) := + match e with + | LCvar _ _ v => fun _ => otrue1 + | LCfield _ x _ _ _ pfieldi e' => fun wf => + synth_lexpr_ocond tmp e' wf + | LCindex _ _ _ _ _ e_arr e_idx => fun wf => + synth_lexpr_ocond tmp e_arr (car wf) /\ + synth_expr_ocond tmp e_idx (cdr wf) /\ + {{ fun se => 0 <= synth_expr_spec tmp e_idx (cdr wf) se < Hindex_size }}%Z + | LChash _ _ _ _ _ e_arr e_idx => fun wf => + synth_lexpr_ocond tmp e_arr (car wf) /\ + synth_expr_ocond tmp e_idx (cdr wf) + end%oprop1. + + (* The only passthrough-related condition that has to be proved manually by the programmer + is for the respec construct. *) + Fixpoint synth_expr_passthrough_ocond + (tmp : AList.t hyper_type_pair) + `{HyperTypeImpl} + (e : expr_constr tp) + : synth_expr_wellformed tmp e -> OProp := + match e with + | ECconst_int _ _ f i => fun _ => otrue + | ECconst_int256 _ _ f i => fun _ => otrue + | ECtempvar _ _ t => fun _ => otrue + | ECbuiltin0 _ _ _ => fun _ => otrue + | ECbuiltin1 _ _ _ _ _ e' => fun wf => synth_expr_passthrough_ocond tmp e' wf + | ECunop _ op _ _ _ _ e' => fun wf => synth_expr_passthrough_ocond tmp e' wf + | ECbinop _ op _ _ _ _ _ _ el er => fun wf => + synth_expr_passthrough_ocond tmp el (car wf) /\ + synth_expr_passthrough_ocond tmp er (cdr wf) + end%oprop + + with synth_lexpr_passthrough_ocond tmp `{HyperTypeImpl}(e : lexpr_constr tp) + : synth_lexpr_wellformed tmp e -> OProp := + match e with + | LCvar _ _ v => fun _ => otrue + | LCfield _ x _ _ _ pfieldi e' => fun wf => + synth_lexpr_passthrough_ocond tmp e' wf + | LCindex _ _ _ _ _ e_arr e_idx => fun wf => + synth_lexpr_passthrough_ocond tmp e_arr (car wf) /\ + synth_expr_passthrough_ocond tmp e_idx (cdr wf) + | LChash _ _ _ _ _ e_arr e_idx => fun wf => + synth_lexpr_passthrough_ocond tmp e_arr (car wf) /\ + synth_expr_passthrough_ocond tmp e_idx (cdr wf) + end%oprop. + + (* Correctness proofs for expression synthesis! *) + Lemma synth_expr_expr_preserves_type tmp `{HyperTypeImpl} e : + synth_expr_wellformed tmp e -> + typeof (synth_expr_expr tmp e) = unpair_ty tp. + (* typeof takes a Clight expression and computes its type. *) + Proof. + generalize dependent tmp. + induction e; try reflexivity. + Qed. + + Lemma synth_lexpr_expr_preserves_type tmp `{HyperTypeImpl} e : + (* synth_lexpr_wellformed tmp e -> *) + typeof (synth_lexpr_expr tmp e) = unpair_ty tp. + Proof. + destruct e; reflexivity. + Qed. + + (* Generally speaking, we tend to pass around just values of type ft, rather than + a dependent type of the value and a proof, and then we prove the ft_cond separately. *) + Theorem synth_expr_spec_satisfies_ft_cond {tmp}`{HyperType} e : + forall wf se, expr_constr_prf e -> + oProp1 (synth_expr_ocond tmp e wf) se -> senv_cond se -> + ht_ft_cond (synth_expr_spec tmp e wf se). + Proof. + intros wf se ec c se_cond. induction ec; simpl in wf, c |- *. + - (* ECconst_int *) + apply fc. + - (* ECtempvar *) + apply se_cond. + - (* ECbuiltin0 *) + apply Hbuiltin0_returns. + - (* ECbuiltin1 *) + apply Hbuiltin1_returns. + apply IHec. apply c. + - (* ECunop *) + rewrite oand1_distr, OProp1map1, <- Hunary_ocond_same in c; [| exact I ]. + apply Hunary_returns; [ eapply IHec |]; + solve [ (* apply le_cond | *) apply c ]. + - (* ECbinop *) + rewrite 2 oand1_distr, OProp2map1, <- Hbinary_ocond_same in c; try exact I. + apply Hbinary_returns; [ eapply IHec1 | eapply IHec2 |]; + solve [ (* apply le_cond | *) apply c ]. + Qed. + + (* Should move this lemma to Integers.v *) + Lemma Int256_lt_repr : forall x y, + (0 <= x < Int256.modulus)%Z -> + (0 <= y < Int256.modulus)%Z -> + (Int256.lt (Int256.repr x) (Int256.repr y) = true) <-> (x < y)%Z. + Admitted. + +Require Import Omega. + + (* This is the main theorem, relating the C expression and the desugaring. *) + Theorem synth_expr_spec_correct {tmp}`{HyperType} e : + forall wf se le, expr_constr_prf e -> + oProp1 (synth_expr_ocond tmp e wf) se -> lenv_cond se le -> + exists v, + (forall m, eval_rvalue ctx me m le (synth_expr_expr tmp e) v) /\ + ht_rel (synth_expr_spec tmp e wf se) v + + with synth_lexpr_spec_correct {tmp}`{HyperType} e : + forall wf se le, lexpr_constr_prf e -> + oProp1 (synth_lexpr_ocond tmp e wf) se -> lenv_cond se le -> + let ltp := synth_lexpr_spec tmp e wf se in (*= (i, ofs) -> *) + (forall m, + eval_lvalue ctx me m le (synth_lexpr_expr tmp e) ltp.(ltype_ident)) /\ + HyperLType ltp. + Proof. + { (* synth_expr_spec_correct *) + intros wf se le ec c le_cond (* gcond *); induction ec; simpl in * |- *. + + - (* ECconst_int256 *) + exists (Vint i); split; [ constructor |]. + + unfold ht_rel; rewrite rel. + constructor; reflexivity. + + - (* ECtempvar *) + assert (some_rel := proj2 (le_cond _ _ _ wf)). + inversion some_rel as [| ? v rel f_eq le_t_eq ]; subst. + exists v; split; [ constructor; symmetry |]; assumption. + + - (* ECbuiltin0 *) + exists (me_query me (Qcall0 Hbuiltin0)). + split. + + intros m. + apply eval_Ecall0. + reflexivity. + + unfold ht_rel. + rewrite Hbuiltin0_correct. + apply CVMval. + + - (* ECbuiltin1 *) + destruct (IHec wf c) as (v & eval & rel). + assert (p_cond := synth_expr_spec_satisfies_ft_cond e + wf _ ec c (lenv_senv_cond le_cond)). + exists (me_query me (Qcall1 Hbuiltin1 v)). + split. + + intros m. + apply eval_Ecall1 with v; only 1: apply eval. + reflexivity. + + unfold ht_rel. + assert (builtin1_correct := Hbuiltin1_correct me _ _ p_cond eq_refl). + destruct builtin1_correct as (v' & Harg & Hret). + rewrite Hret. + replace v with v' by (inversion rel; congruence). + apply CVMval. + + - (* ECunop *) + rewrite oand1_distr, OProp1map1, <- Hunary_ocond_same in c; [| exact I ]. + destruct (IHec wf (proj1 c)) as (v & eval & rel). + assert (p_cond := synth_expr_spec_satisfies_ft_cond e + wf _ ec (proj1 c) (lenv_senv_cond le_cond)). + assert (unary_correct := Hunary_correct _ _ p_cond (proj2 c) eq_refl). + inversion unary_correct as [| ? ? ? f_eq cval_op_eq ]; subst. + eapply eq_sym, cval_sem_unary_operation in cval_op_eq; [| exact rel ]. + destruct cval_op_eq as (v' & sem_eq & rel'). + + exists v'; split; [| assumption ]. + intros m. + eapply eval_Eunop with (v1:=v); [ apply eval |]. + apply sem_eq. + + - (* ECbinop *) + rewrite 2 oand1_distr, OProp2map1, <- Hbinary_ocond_same in c; try exact I. + destruct c as (c1 & c2 & binary_cond). + destruct (IHec1 (car wf) c1) as (v1 & eval1 & rel1), + (IHec2 (cdr wf) c2) as (v2 & eval2 & rel2). + assert (p1_cond := synth_expr_spec_satisfies_ft_cond el (car wf) _ ec1 c1 (lenv_senv_cond le_cond)). + assert (p2_cond := synth_expr_spec_satisfies_ft_cond er (cdr wf) _ ec2 c2 (lenv_senv_cond le_cond)). + assert (bin_correct := Hbinary_correct _ _ _ _ p1_cond p2_cond + binary_cond eq_refl eq_refl). + inversion bin_correct as [| ? ? ? f_eq cval_op_eq ]; subst. + eapply eq_sym, cval_sem_binary_operation in cval_op_eq; + [| exact rel1 | exact rel2 ]. + destruct cval_op_eq as (v' & sem_eq & rel'). + + exists v'; split; [| assumption ]. + intros m. + eapply eval_Ebinop with _ _ v1 v2; [ apply eval1 | apply eval2 |]. + apply sem_eq. + } + + { (* synth_lexpr_spec_ptr_correct *) + intros wf se le ec c le_cond (* gcond *). + induction ec; intros; simpl in * |- *. + - (* LCvar *) + unfold ltp. + assert (v_ltype := v_prf.(VARltype)). + assert (absorption := fun A B b f => @conj A B (f b) b); + apply absorption; clear absorption. + + (* HyperLtype *) + exact v_ltype. + + (* eval_lvalue *) + intros hyper_ltype m. + rewrite v_prf.(VARoffset_zero). + simpl. + apply eval_Eglob. + reflexivity. + - (* LCfield *) + destruct (IHec _ c) as (eval & struct_ltype). + unfold ltp; split; [| eauto with typeclass_instances ]. + + destruct Hfield_delta_correct + as (tid & flds & tpl_eq (* & field_type_eq & field_offset_eq *)). + intros m. + eapply eval_Efield. + + apply eval. + + unfold ExtEnv.access_field. + rewrite synth_lexpr_expr_preserves_type. + destruct tpl_eq as [fList [tpl_eq1 tpl_eq2]]. + rewrite tpl_eq1. + rewrite tpl_eq2. + reflexivity. + - (* LCindex *) + rewrite 2 oand1_distr in c. + destruct c as (arr_c & idx_c & idx_in_bound). + destruct (IHec _ arr_c) as (array_eval & array_ltype). + destruct (synth_expr_spec_correct _ _ _ _ _ _ _ _ + H0 idx_c le_cond) as (index & index_eval & index_rel). + + unfold ltp; split; [| eauto with typeclass_instances ]. + intros m. + eapply eval_Eindex. + + (* eval_expr array *) + simpl. + apply array_eval. + + inversion index_rel. + subst index v. + apply index_eval. + + simpl. + inversion index_rel. + subst index v. + unfold ExtEnv.index_array. + rewrite (synth_lexpr_expr_preserves_type tmp e). + rewrite Hindex_array_type. + simpl in idx_in_bound. + destruct idx_in_bound as [idx_in_bound1 idx_in_bound2]. + assert (0 <= synth_expr_spec tmp idx (cdr wf) se < Int256.modulus)%Z + as idx_in_bound_modulus. + { + split. + - exact idx_in_bound1. + - transitivity Hindex_size. + exact idx_in_bound2. + destruct Hindex_size_bound; assumption. + } + match goal with [|- (if ?X then _ else _) = _] => + replace X with true + by (symmetry; + apply (proj2 (Int256_lt_repr _ _ idx_in_bound_modulus Hindex_size_bound) idx_in_bound2)) + end. + reflexivity. + - (* LChash *) + rewrite oand1_distr in c. + destruct c as (arr_c & idx_c). + destruct (IHec _ arr_c) as (array_eval & array_ltype). + destruct (synth_expr_spec_correct _ _ _ _ _ _ _ _ + H0 idx_c le_cond) as (index & index_eval & index_rel). + + unfold ltp; split; [| eauto with typeclass_instances ]. + + intros m. + eapply eval_Eindex. + + (* eval_hash *) + simpl. + apply array_eval. + + inversion index_rel. + subst index v. + apply index_eval. + + simpl. + inversion index_rel. + subst index v. + unfold ExtEnv.index_array. + rewrite (synth_lexpr_expr_preserves_type tmp e). + rewrite Hhash_type. + reflexivity. + } + Qed. + + Theorem synth_lexpr_spec_is_ghost_eq {tmp}`{HyperTypeImpl} e wf se : + (synth_lexpr_spec tmp e wf se).(ltype_ghost) = lexpr_is_ghost e. + Proof. + induction e; solve [ reflexivity | apply IHe ]. + Qed. + + Lemma synth_lexpr_passthrough tmp `{HyperTypeImpl} e wf : + lexpr_constr_passthrough_prf e -> + oProp (synth_lexpr_passthrough_ocond tmp e wf) -> + forall se, variable_passthrough_prf (synth_lexpr_spec tmp e wf se). + Proof. + Admitted. + +End EXPR_FUNC. diff --git a/src/core/SynthesisFunc.v b/src/core/SynthesisFunc.v new file mode 100755 index 0000000..5401914 --- /dev/null +++ b/src/core/SynthesisFunc.v @@ -0,0 +1,269 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** Synthesizing Clight functions and specifications from DeepSpec along + with the proofs. *) + +(* Standard library modules *) +Require Import BinPos. +Require Import List. + +(* CompCert modules *) +Require Import backend.AST. +(*Require Import compcert.common.Events. +Require Import compcert.common.Globalenvs. *) +Require Import backend.Values.HighValues. +Require Import backend.MemoryModel. +Require Import cclib.Coqlib. +Require Import cclib.Maps. +Require Import backend.Ctypes. +Require Import backend.Expressions.ExpMiniC. +Require Import backend.Statements.StmtMiniC. +Require Import backend.MachineModel. + +(* CompCertX modules *) +(* Require Import compcertx.cfrontend.ClightBigstepX. *) + +(* DeepSpec modules *) +Require Import DeepSpec.lib.IndexedMaps. +Require Import DeepSpec.lib.SimpleMaps. +Require Import DeepSpec.lib.SimpleIndexedMaps. +Require Import DeepSpec.lib.OProp. +Require DeepSpec.lib.LangDef. +Require Import DeepSpec.core.SEnv. +Require Import DeepSpec.core.Cval. +Require Import DeepSpec.core.HyperType. +(*Require Import DeepSpec.core.HyperMem.*) +Require Import DeepSpec.core.MemoryModel. +Require Import DeepSpec.core.Syntax. +Require Import DeepSpec.core.SynthesisStmt. + +Section FUNC_FUNC. + Open Scope positive_scope. + + Context`{TheLayerSpec : LayerSpecClass}. (* TODO: this line will go away when we fix the type of HyperMem. *) + Context`{HM : HyperMem}. + + Let genv_t := genv. + + Definition func_spec_record returns := + (GetHighData * (* trace * *) tp_ft returns)%type. + Definition mk_func_spec_record returns m (*t*) f : func_spec_record returns := + ((m (*, t *)), f). + Definition sf_mem {returns}(r : func_spec_record returns) := (*fst*) (fst r). + (*Definition sf_trace {returns}(r : func_spec_record returns) := snd (fst r).*) + Definition sf_return {returns}(r : func_spec_record returns) := snd r. + + (* Statement desugaring returns a record, here we translate that into a func_spec_record (which is actually a pair, because the mcertikos framework didn't use a record here).*) + Definition stmt_to_func_spec_record {tmp returns} + (s : spec_env_t tmp -> GetHighData -> stmt_spec_record returns) + se m + := let s' := s se m + in mk_func_spec_record returns s'.(ss_mem) (*s'.(ss_trace)*) s'.(ss_return). + + Record func_output_equivalent {returns} (j:meminj) (ml:mem) v + (r : func_spec_record returns) : Prop := mk_func_output_equivalent{ +(* fo_mem : mem_match j (sf_mem r) ml; *) + fo_return : function_return_dec returns = true -> cval_match v (ht_cval (sf_return r)); + fo_ft_cond : ht_ft_cond (sf_return r) + }. +(* Global Arguments fo_mem {_ _ _ _ _} _. *) + Global Arguments fo_return {_ _ _ _ _} _ _. + Global Arguments fo_ft_cond {_ _ _ _ _} _. + + Definition synth_func_pure f := synth_stmt_pure f.(FC_body). + + (* We use fancy dependent types for hetrogeneously typed lists of arguments, and variable-arity functions. *) + Definition param_folder p (htp : hyper_type_pair) + := (Pos.succ (fst p), AList.set (fst p) htp (snd p)). + Definition param_env params ident_start + := snd (fold_left param_folder params (ident_start, AList.empty)). + Fixpoint param_func_type (params : list hyper_type_pair) T : Type + := match params with + | nil => T + | htp :: l => tp_ft htp -> param_func_type l T + end. + + Fixpoint abs_param_func' params T : + forall ident_start tmp, SpecTree.t tmp -> + (SpecTree.t (snd (fold_left param_folder params (ident_start, tmp))) + -> T) -> + param_func_type params T := + match params as p + return forall ident_start tmp, SpecTree.t tmp -> + (SpecTree.t (snd (fold_left param_folder p (ident_start, tmp))) -> T) -> + param_func_type p T with + | nil => fun _ _ se t => t se + | htp :: l => fun ident_start tmp se t => + fun f => abs_param_func' l T + (Pos.succ ident_start) (AList.set ident_start htp tmp) + (SpecTree.set ident_start htp f se) t + end. + Definition abs_param_func {params T} ident_start : + (SpecTree.t (param_env params ident_start) -> T) -> + param_func_type params T := + abs_param_func' params T ident_start (AList.empty) SpecTree.empty. + + Fixpoint hlist_param_func' params T (hlist : HList tp_ft params) : + forall ident_start tmp, SpecTree.t tmp -> + (SpecTree.t (snd (fold_left param_folder params (ident_start, tmp))) + -> T) -> T := + match hlist in (HList _ p) + return forall ident_start tmp, SpecTree.t tmp -> + (SpecTree.t (snd (fold_left param_folder p (ident_start, tmp))) + -> T) -> T with + | HNil => fun _ _ se t => t se + | HCons htp l f rest => fun ident_start tmp se t => + hlist_param_func' l T rest + (Pos.succ ident_start) (AList.set ident_start htp tmp) + (SpecTree.set ident_start htp f se) t + end. + Definition hlist_param_func {params T} ident_start + (t : SpecTree.t (param_env params ident_start) -> T) + (hlist : HList tp_ft params) : T + := hlist_param_func' params T hlist ident_start _ SpecTree.empty t. + + Definition fn_params_folder p (htp : hyper_type_pair) + := (Pos.succ (fst p), (fst p, tp_ty htp) :: (snd p)). + Definition fn_params_fold params ident_start + := rev' (snd (fold_left fn_params_folder params (ident_start, nil))). + + Definition synth_func_wellformed (f : function_constr) : Type := + let param_tmp := param_env f.(FC_params) f.(FC_param_ident_start) in + synth_stmt_wellformed f.(FC_body) f.(FC_ident_start) param_tmp. + + (* Naming convention, synth_func_ are the synthesis functions returning various things, + in this case returing a Clight function. *) + Definition synth_func_func (f : function_constr) : function := + let dest := f.(FC_ident_start) in + let ret_ty := tp_ty f.(FC_returns) in + let param_tmp := param_env f.(FC_params) f.(FC_param_ident_start) in + let body := synth_stmt_stmt f.(FC_body) dest param_tmp in {| + fn_return := ret_ty; + (*fn_callconv := cc_default;*) + fn_params := fn_params_fold f.(FC_params) f.(FC_param_ident_start); + fn_temps := (dest, ret_ty) :: synth_stmt_locals f.(FC_body) dest param_tmp; + fn_locals := nil; + fn_body := if function_return_dec f.(FC_returns) + then Ssequence body (Sreturn (Some xH)) + else Ssequence body (Sreturn None) + |}. + + (* return just the return value *) + Definition synth_func_spec_ret (f : function_constr) + (wf : synth_func_wellformed f) + : param_func_type f.(FC_params) (machine_env GetHighData -> GetHighData -> tp_ft f.(FC_returns)) := + let param_tmp := param_env f.(FC_params) f.(FC_param_ident_start) in + abs_param_func f.(FC_param_ident_start) + (fun se me m => synth_stmt_spec_ret me f.(FC_body) f.(FC_ident_start) + param_tmp wf se m). +(* + (* return both return value and new memory. *) + Definition synth_func_spec (f : function_constr) + (wf : synth_func_wellformed f) + : param_func_type f.(FC_params) (GetHighData -> GetHighData * tp_ft f.(FC_returns)) := + let param_tmp := param_env f.(FC_params) f.(FC_param_ident_start) in + abs_param_func f.(FC_param_ident_start) + (fun se m => synth_stmt_spec_cont f.(FC_body) f.(FC_ident_start) + param_tmp wf se m + _ (fun r => (r.(ss_mem), r.(ss_return)))). +*) + (* return both return value and new memory, monadic version. *) + Definition synth_func_spec_opt (f : function_constr) + (wf : synth_func_wellformed f) + (* r (k : func_spec_record f.(FC_returns) -> option r) *) + : param_func_type f.(FC_params) + (machine_env GetHighData -> DS (tp_ft (FC_returns f))) := + (* (GetHighData -> option r) := *) + let param_tmp := param_env f.(FC_params) f.(FC_param_ident_start) in + abs_param_func f.(FC_param_ident_start) + (fun se me => synth_stmt_spec_opt me f.(FC_body) f.(FC_ident_start) + param_tmp wf se). + + (* TODO: May need [ht_ft_cond] of the arguments for VC dischanging to go through. + In the part of mcertikos that we have reimplemented so far this didn't happen because all ints + were loop variables (so they automatically have bounded ranges) but that will not hold in general. *) + Definition synth_func_cond (f : function_constr)(wf : synth_func_wellformed f) := + let param_tmp := param_env f.(FC_params) f.(FC_param_ident_start) in + abs_param_func f.(FC_param_ident_start) (fun se me m => + (*fix: *) (* senv_cond se -> *) + if function_return_dec f.(FC_returns) then + synth_stmt_veri_cond me f.(FC_body) f.(FC_ident_start) param_tmp wf se m /\ + ht_valid_ft_cond + (synth_stmt_spec_ret me f.(FC_body) f.(FC_ident_start) + param_tmp wf se m) + else + synth_stmt_veri_cond me f.(FC_body) f.(FC_ident_start) param_tmp wf se m + ). + + (* TODO: this might also need a senv_cond condition, don't know. (There is a curried se somewhere in here.) *) + Definition synth_func_obligation (f : function_constr)(wf : synth_func_wellformed f) := + let param_tmp := param_env f.(FC_params) f.(FC_param_ident_start) in + abs_param_func f.(FC_param_ident_start) + (fun se me d => synth_stmt_obligation me f.(FC_body) f.(FC_ident_start) param_tmp wf se d). + + Fixpoint spec_environment (htps : list hyper_type_pair) + : forall (args : HList tp_ft htps) i (rest_typ : AList.t hyper_type_pair) (rest : SpecTree.t rest_typ), SpecTree.t (snd (fold_left param_folder htps (i, rest_typ))) := + match htps + return forall (args : HList tp_ft htps) i (rest_typ : AList.t hyper_type_pair) (rest : SpecTree.t rest_typ), SpecTree.t (snd (fold_left param_folder htps (i, rest_typ))) + with + | nil => fun args ident rest_typ rest => rest + | htp::htps => fun args i rest_typ rest => + spec_environment htps (hlist_tl args) (Pos.succ i) (AList.set i htp rest_typ) (SpecTree.set i _ (hlist_hd args) rest) + end. + + Fixpoint apply_param_func {params T} : + param_func_type params T -> HList tp_ft params -> T := + match params with + | nil => fun t _ => t + | htp :: l => fun f hlist => @apply_param_func l T + (f (hlist_hd hlist)) (hlist_tl hlist) + end. + + Lemma apply_abs' + : forall T htps i tmp (se : SpecTree.t tmp) + (f : SpecTree.t (snd (fold_left param_folder htps (i, tmp))) -> T) + args, + apply_param_func (abs_param_func' htps T i tmp se f) args + = f (spec_environment htps args i _ se). + Proof. + induction htps; intros. + + reflexivity. + + simpl. + rewrite IHhtps. + reflexivity. + Qed. + + Lemma apply_abs + : forall htps T i + (f : SpecTree.t (param_env htps i) -> T) + args, + apply_param_func (abs_param_func i f) args + = f (spec_environment htps args i _ (SpecTree.empty)). + Proof. + intros. + apply apply_abs'. + Qed. + +End FUNC_FUNC. diff --git a/src/core/SynthesisStmt.v b/src/core/SynthesisStmt.v new file mode 100755 index 0000000..9c8c227 --- /dev/null +++ b/src/core/SynthesisStmt.v @@ -0,0 +1,2981 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** Synthesizing Clight statements, functional statements, and + specifications from DeepSpec along with the proofs. *) +(* Standard library modules *) +Require Import BinIntDef. +Require Import BinPosDef. +Require Import Bool. + +(* CompCert modules *) +Require Import backend.AST. +(*Require Import compcert.common.Events. +Require Import compcert.common.Globalenvs. (* Genv *) *) +Require Import backend.MemoryModel. +Require Import backend.Values.HighValues. +Require Import cclib.Coqlib. +Require Import cclib.Integers. +Require Import cclib.Maps. (* PTree *) +Require Import backend.Expressions.ExpMiniC. +Require Import backend.Statements.StmtMiniC. +(* Require Import compcert.cfrontend.ClightBigstep. *) +Require Import backend.Cop. (* bool_val *) +Require Import backend.Ctypes. +Require Import backend.MachineModel. + +(* DeepSpec modules *) +Require Import DeepSpec.lib.OProp. +(* Require Import DeepSpec.lib.Monad.ContOptionMonad. *) +Require Import DeepSpec.core.Cval. +Require Import DeepSpec.core.SEnv. +Require Import DeepSpec.core.HyperType. +Require Import DeepSpec.core.HyperTypeInst. +(*Require Import DeepSpec.core.HyperMem.*) +Require Import DeepSpec.core.MemoryModel. +Require Import DeepSpec.core.Syntax. +Require Import DeepSpec.core.SynthesisExpr. +Require DeepSpec.lib.LangDef. + +Require Import DeepSpec.lib.SimpleMaps. +Require Import DeepSpec.lib.SimpleIndexedMaps. +Require Import backend.phase.MiniC.Semantics. +Require Import backend.phase.MiniC.BigstepSemantics. +(* Require Import backend.Events. *) + +Local Open Scope Z. +Opaque Monad.bind Monad.ret (* OptionMonad.guard *) MonadState.get MonadState.put MonadState.gets MonadState.modify StateMonad.runStateT StateMonad.execStateT StateMonad.evalStateT MonadZero.mzero. + +Section STMT_FUNC. + Context`{HM : HyperMem}. + Context (me : machine_env GetHighData). + + (*Variable m0 : MemLow.*) + + (* This is the type of what the desugaring returns: new abstract data, and a return value. *) + Record stmt_spec_record returns : Type + := mk_stmt_spec_record{ + ss_mem : GetHighData; + (*ss_trace : trace;*) + ss_return : tp_ft returns (* If the DeepSpec function does not return a value, the specification returns unit. *) + }. + (* begin hide *) + Global Arguments mk_stmt_spec_record _ _ (*_*) _. + Global Arguments ss_mem {_} _. + (* Global Arguments ss_trace {_} _. *) + Global Arguments ss_return {_} _. + (* end hide *) + + Lemma stmt_spec_record_surjective {returns} r : + r = mk_stmt_spec_record returns r.(ss_mem) (*r.(ss_trace)*) r.(ss_return). + Proof. + destruct r; reflexivity. + Qed. + + (* This is the output of the correctness theorem. *) + Record stmt_output_equivalent {tmp returns} + (j : meminj) (* memory injection, we thread this through because we use it for mem_match. *) + dest (* identifier for the temporary variable used to store the return value of the statement. *) + (ml : MemLow) (* low memory after the end of statement execution.. *) + (se : spec_env_t tmp) (* This is the specification environment (immutable, so same both before and after statement execution) *) + le (* Temporary environment after the execution. *) + (r : stmt_spec_record returns) : Prop := mk_stmt_output_equivalent{ + so_mem : mem_match j r.(ss_mem) ml; (* The resulting low memory matches the resulting abstract data. *) + so_se : lenv_cond se le; (* The values in le all refine the values in the se. *) + so_return : function_return_dec returns = true -> ht_rel_some r.(ss_return) (le ! dest); (* If the function returns a value, then it is stored in dest. *) + so_ft_cond : ht_ft_cond r.(ss_return) + }. + (* begin hide *) + Global Arguments so_mem {_ _ _ _ _ _ _ _} _. + Global Arguments so_se {_ _ _ _ _ _ _ _} _ _ _ _ _. + Global Arguments so_return {_ _ _ _ _ _ _ _} _ _. + Global Arguments so_ft_cond {_ _ _ _ _ _ _ _} _. + (* end hide *) + + + (* Auxillary functions about isomorphisms. These are used for looking things up in the Indexed Trees. *) + Definition gss_iso {tmp i}{htp : hyper_type_pair} : + ((AList.set i htp tmp) ! i ~~~ Some htp)%alist. + Proof. + simpl. + rewrite Pos.eqb_refl. + exact identity_automorphism. + Defined. + + Definition gso_iso {tmp i j}{htp htp' : hyper_type_pair} : + i <> j -> + (tmp ! i ~~~ Some htp)%alist -> + ((AList.set j htp' tmp) ! i ~~~ Some htp)%alist. + Proof. + intros. + rewrite AList.gso by assumption. + exact X. + Defined. + + (* These are auxilliary definitions used for desugaring loops. (I.e. the desugared function will contain calls to these.) *) + Definition for_start m (start_idx : Z) := + (@mk_stmt_spec_record void_unit_pair m (*E0*) tt, start_idx). + Definition for_step {tmp id_it}(se : spec_env_t tmp) : + let tmp' := AList.set id_it int_Z32_pair tmp in + (SpecTree.t tmp' -> GetHighData -> stmt_spec_record void_unit_pair) -> + stmt_spec_record void_unit_pair * Z -> + stmt_spec_record void_unit_pair * Z := + fun f p => + let '(r, idx) := p in + let r' := f (SpecTree.set id_it int_Z32_pair idx se) r.(ss_mem) in + ({| + ss_mem := r'.(ss_mem); + (*ss_trace := r.(ss_trace) ** r'.(ss_trace);*) + ss_return := r'.(ss_return) + |}, idx + 1). + + Definition fold_start htp (start_idx : Z)(start_recur : tp_ft htp) := + (start_recur, start_idx). + Definition fold_step {htp tmp id_it id_recur}(se : spec_env_t tmp) m : + let tmp' := AList.set id_recur htp (AList.set id_it int_Z32_pair tmp) in + (SpecTree.t tmp' -> GetHighData -> stmt_spec_record htp) -> + tp_ft htp * Z -> tp_ft htp * Z := + fun f p => + let '(r, idx) := p in + ((f (SpecTree.set id_recur htp r + (SpecTree.set id_it int_Z32_pair idx se)) m).(ss_return), + idx + 1). + + Definition dest_not_exist dest (tmp : AList.t hyper_type_pair) := + (tmp ! dest ~~~ None)%alist. + Lemma dest_not_exist_ext {tmp dest} id htp : + dest_not_exist dest (AList.set id htp tmp) -> + dest_not_exist dest tmp. + Proof. + unfold dest_not_exist. + intros dest_not_exist_set. + destruct (Pos.eq_dec dest id) as [ dest_eq | dest_ne ]. + - exfalso. + rewrite dest_eq in dest_not_exist_set. + rewrite AList.gss in dest_not_exist_set. + set (is_some (o : option hyper_type_pair) := + match o with Some _ => true | _ => false end). + apply (isomorphism_decidable_neq_absurd is_some dest_not_exist_set); + reflexivity. + - rewrite AList.gso in dest_not_exist_set; assumption. + Qed. + + Definition yield_expr dest e := + Sset dest e. + + (* This is used as hypothesis for some of the later theorems. + Compare with figure 4.9 "Command synthesis" in the thesis: + some elaboration rule has a precondition that e.g. the test + expression in a loop has to be pure. If it's not, we can't + synthesize, or at least not prove any theorems about the code. *) + Fixpoint synth_stmt_pure {returns}(c : cmd_constr returns) : bool := + match c with + | CCskip => true + | CClet _ tp _ id c1 c2 => synth_stmt_pure c1 && synth_stmt_pure c2 + | CCload tp _ e => true + | CCstore _ _ el er => false + | CCsequence _ c1 c2 => synth_stmt_pure c1 && synth_stmt_pure c2 + | CCifthenelse _ e c_true c_false => + synth_stmt_pure c_true && synth_stmt_pure c_false + | CCfor id_it id_end e1 e2 c3 => false + (* for loops are assumed to always have side effects. (otherwise, why use the loop?) *) + | CCfirst _ id_it id_end id_dest e1 e2 c3 c4 c5 => + synth_stmt_pure c3 && synth_stmt_pure c4 && synth_stmt_pure c5 + | CCfold tp _ id_it id_end id_recur id_dest e1 e2 e3 c4 => true + | CCcall argt ret prim args => prim.(PRIMpure) + (* | CCcall_ext argt ret addr prim args => prim.(PRIMpure) *) + | CCyield tp _ e => true + | CCconstr _ _ _ _ _ _ _ => false + | CCassert c => synth_stmt_pure c + (* | CCghost c => synth_stmt_pure c *) + | CCdeny c => synth_stmt_pure c + | CCpanic _ _ => true + | CCrespec _ _ c _ => true + | CCrespec_opt _ _ c _ => false + end. + + + (* More typechecking judgements. This requires + -again, that temporaries are welltyped, see e.g. the case for CClet. + -and the pure conditions, see e.g. CCfold. + *) + Definition fold_synth_expr_wellformed tmp {tps}(es : expr_constr_list tps) := + HList_fold_right_nodep + (fun htp (r : expr_constr _) a => prod (synth_expr_wellformed tmp (tp := tp_type_pair htp) r) a) + True + es. + + Fixpoint synth_stmt_wellformed {returns}(c : cmd_constr returns) dest tmp : Type := + match c with + | CCskip => dest_not_exist dest tmp + | CClet _ tp _ id c1 c2 => + let tmp' := AList.set id (mk_hyper_type_pair tp) tmp in + ((tmp ! id ~~~ None) * dest_not_exist dest tmp') * + (synth_stmt_wellformed c1 id tmp * synth_stmt_wellformed c2 dest tmp') + | CCload tp _ e => synth_lexpr_wellformed tmp e + | CCstore _ _ el er => + synth_lexpr_wellformed tmp el * synth_expr_wellformed tmp er + | CCsequence _ c1 c2 => + synth_stmt_wellformed c1 dest tmp * synth_stmt_wellformed c2 dest tmp + | CCifthenelse _ e c_true c_false => + synth_expr_wellformed tmp e * + (synth_stmt_wellformed c_true dest tmp * + synth_stmt_wellformed c_false dest tmp) + | CCfor id_it id_end e1 e2 c3 => + let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + dest_not_exist dest tmp'')) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + synth_stmt_wellformed c3 dest tmp'')) + | CCfirst _ id_it id_end id_dest e1 e2 c3 c4 c5 => + let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + ((tmp'' ! id_dest ~~~ None) * (dest_not_exist dest tmp'' * + ((synth_stmt_pure c3 ~~~ true)))))) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + (synth_stmt_wellformed c3 id_dest tmp'' * + (synth_stmt_wellformed c4 dest tmp'' * + synth_stmt_wellformed c5 dest tmp'')))) + | CCfold tp _ id_it id_end id_recur id_dest e1 e2 e3 c4 => + let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + let tmp''' := AList.set id_recur (mk_hyper_type_pair tp) tmp'' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + ((tmp'' ! id_recur ~~~ None) * ((tmp''' ! id_dest ~~~ None) * (* we use ~~~ because it has to be computable, we will use it as argument to the tree get function. *) + (dest_not_exist dest tmp''' * (synth_stmt_pure c4 ~~~ true)))))) * (* we could probably use = instead of ~~~ in this one case. *) + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + (synth_expr_wellformed tmp e3 * synth_stmt_wellformed c4 id_dest tmp'''))) + | CCcall argt ret prim args => fold_synth_expr_wellformed tmp args + (* | CCcall_ext argt ret addr prim args => fold_synth_expr_wellformed tmp args *) + | CCconstr _ _ _ _ el flds _ => + synth_lexpr_wellformed tmp el * fold_synth_expr_wellformed tmp flds + + | CCyield tp _ e => synth_expr_wellformed tmp e + | CCassert c => (synth_stmt_pure c ~~~ true) * synth_stmt_wellformed c dest tmp + | CCdeny c => (synth_stmt_pure c ~~~ true) * synth_stmt_wellformed c dest tmp + (* | CCghost c => synth_stmt_wellformed c dest tmp *) + | CCpanic _ _ => True + | CCrespec _ tmp' c spec => + ((tmp ~~~ tmp') * (synth_stmt_pure c ~~~ true)) * + synth_stmt_wellformed c dest tmp + | CCrespec_opt _ tmp' c spec => (tmp ~~~ tmp') * synth_stmt_wellformed c dest tmp + end%type%alist. + + (* Collect all the local variables the synthesized C statment will contain. + This is used in SynthesisFunction, we need to give a list of all the + locals of the function body. *) + Fixpoint synth_stmt_locals {returns}(c : cmd_constr returns) dest tmp : + list (ident * type) := + match c with + | CCskip => nil + | CClet _ tp _ id c1 c2 => + (id, unpair_ty tp) :: + synth_stmt_locals c1 id tmp ++ + synth_stmt_locals c2 dest (AList.set id (mk_hyper_type_pair tp) tmp) + | CCload tp _ e => nil + | CCstore _ _ el er => nil + | CCsequence _ c1 c2 => + synth_stmt_locals c1 dest tmp ++ synth_stmt_locals c2 dest tmp + | CCifthenelse _ e c_true c_false => + synth_stmt_locals c_true dest tmp ++ synth_stmt_locals c_false dest tmp + | CCfor id_it id_end e1 e2 c3 => + let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + (id_it, tint) :: (id_end, tint) :: synth_stmt_locals c3 dest tmp'' + | CCfirst _ id_it id_end id_dest e1 e2 c3 c4 c5 => + let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + (id_it, tint) :: (id_end, tint) :: (id_dest, tint) :: + synth_stmt_locals c3 id_dest tmp'' ++ synth_stmt_locals c4 dest tmp'' ++ + synth_stmt_locals c5 dest tmp'' + | CCfold tp _ id_it id_end id_recur id_dest e1 e2 e3 c4 => + let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + let tmp''' := AList.set id_recur (mk_hyper_type_pair tp) tmp'' in + (id_it, tint) :: (id_end, tint) :: (id_recur, unpair_ty tp) :: + (id_dest, unpair_ty tp) :: synth_stmt_locals c4 id_dest tmp''' + | CCcall argt ret prim args => nil + (* | CCcall_ext argt ret addr prim args => nil *) + | CCconstr _ _ _ _ _ _ _ => nil + | CCyield tp _ e => nil + | CCassert _ => nil + | CCdeny _ => nil + (* | CCghost _ => nil *) + | CCpanic _ _ => nil + | CCrespec _ tmp' c cpec => synth_stmt_locals c dest tmp + | CCrespec_opt _ tmp' c spec => synth_stmt_locals c dest tmp + end. + + Fixpoint synth_CCconstr_stmt tmp target fld_ids {fld_tps} + (flds : expr_constr_list fld_tps){struct flds} : statement := + match flds with + | HNil => Sskip + | HCons htp _ fld_e res_es => match fld_ids with + | nil => Sskip (* Impossible, but let's not bother *) + | fld_id :: res_ids => + Ssequence + (Sassign (Efield target fld_id (tp_ty htp)) + (synth_expr_expr tmp fld_e)) + (synth_CCconstr_stmt tmp target res_ids res_es) + end + end. + + (* This builds the Clight statement for a statement (see Figure 4.9). *) + Fixpoint synth_stmt_stmt {returns}(c : cmd_constr returns) dest tmp : statement := + match c with + | CCskip => Sskip + | CClet _ tp _ id c1 c2 => + let tmp' := AList.set id (mk_hyper_type_pair tp) tmp in + Ssequence (synth_stmt_stmt c1 id tmp) (synth_stmt_stmt c2 dest tmp') + | CCload tp _ el => Sset dest (synth_lexpr_expr tmp el) + | CCstore _ _ el er => + if lexpr_is_ghost el then + Sskip + else + Sassign (synth_lexpr_expr tmp el) (synth_expr_expr tmp er) + | CCsequence _ c1 c2 => + Ssequence (synth_stmt_stmt c1 dest tmp) (synth_stmt_stmt c2 dest tmp) + | CCifthenelse _ e c_true c_false => + Sifthenelse (synth_expr_expr tmp e) (synth_stmt_stmt c_true dest tmp) + (synth_stmt_stmt c_false dest tmp) + | CCfor id_it id_end e1 e2 c3 => + let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + Ssequence (Sset id_end (synth_expr_expr tmp e2)) + (Sfor (Sset id_it (synth_expr_expr tmp e1)) + (Ebinop Olt (Etempvar id_it tint) (Etempvar id_end tint) tint) + (synth_stmt_stmt c3 dest tmp'') + (Sset id_it (Ebinop Oadd (Etempvar id_it tint) + (Econst_int256 Int256.one tint) tint))) + | CCfirst _ id_it id_end id_dest e1 e2 c3 c4 c5 => + let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + let s3 := synth_stmt_stmt c3 id_dest tmp'' in + let s4 := synth_stmt_stmt c4 dest tmp'' in + let s5 := synth_stmt_stmt c5 dest tmp'' in + Ssequence + (Sset id_end (synth_expr_expr tmp e2)) + (Ssequence + (Sset id_it (synth_expr_expr tmp e1)) + (Sloop + (Ssequence + (Sifthenelse + (Ebinop Olt (Etempvar id_it tint) (Etempvar id_end tint) tint) + Sskip + (Ssequence + (Sset id_it (Etempvar id_end tint)) + (Ssequence s5 Sbreak))) + (Ssequence + s3 + (Sifthenelse + (Etempvar id_dest tint) + (Ssequence s4 Sbreak) + (Sset id_it (Ebinop Oadd (Etempvar id_it tint) + (Econst_int256 Int256.one tint) tint))))) + )) + | CCfold tp _ id_it id_end id_recur id_dest e1 e2 e3 c4 => + let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + let tmp''' := AList.set id_recur (mk_hyper_type_pair tp) tmp'' in + Ssequence (Ssequence (Sset id_end (synth_expr_expr tmp e2)) + (Sfor (Ssequence (Sset id_it (synth_expr_expr tmp e1)) + (Sset id_recur (synth_expr_expr tmp e3))) + (Ebinop Olt (Etempvar id_it tint) (Etempvar id_end tint) tint) + (synth_stmt_stmt c4 id_dest tmp''') + (Ssequence (Sset id_it (Ebinop Oadd (Etempvar id_it tint) + (Econst_int256 Int256.one tint) tint)) + (Sset id_recur (Etempvar id_dest (unpair_ty tp)))))) + (Sset dest (Etempvar id_recur (unpair_ty tp))) + | CCcall argt ret prim args => + let arg_ty := to_typelist argt in + let prim_type := Tfunction arg_ty (tp_ty ret) (*prim.(PRIMcc)*) in + let synth_expr_htp htp := synth_expr_expr tmp (tp := tp_type_pair htp) in + let return_tmp := if function_return_dec ret then + Some dest + else + None in + Scall return_tmp prim.(PRIMident) + (HList_map_nodep synth_expr_htp args) + | CCyield tp _ e => Sset dest (synth_expr_expr tmp e) + + | CCconstr _ _ fld_ids _ el args _ => + if lexpr_is_ghost el then + Sskip + else + let target := synth_lexpr_expr tmp el in + synth_CCconstr_stmt tmp target fld_ids args + | CCassert c1 => + Ssequence (synth_stmt_stmt c1 dest tmp) + (Sifthenelse (Etempvar dest tint) + Sskip + Srevert) + | CCdeny c1 => + Ssequence (synth_stmt_stmt c1 dest tmp) + (Sifthenelse (Etempvar dest tint) + Srevert + Sskip) + (* | CCghost _ => Sskip *) + | CCpanic _ _ => Srevert + | CCrespec _ tmp' c spec => synth_stmt_stmt c dest tmp + | CCrespec_opt _ tmp' c spec => synth_stmt_stmt c dest tmp + end. + + Fixpoint map2_synth_expr_spec {tmp tps} es (se : spec_env_t tmp) + : @fold_synth_expr_wellformed tmp tps es -> HList tp_ft tps := + match es with + | HNil => fun _ => HNil + | HCons x ls e es => fun wf => + let (wf, IHwf) := wf + in HCons x ls (synth_expr_spec me tmp e wf se) (map2_synth_expr_spec es se IHwf) + end. + +Require Import DeepSpec.lib.Monad.Monad. +Require Import DeepSpec.lib.Monad.MonadState. +Require Import DeepSpec.lib.Monad.MonadZero. +Require Import DeepSpec.lib.Monad.StateMonad. +Require Import DeepSpec.lib.Monad.StateMonadOption. + +Import MonadNotation. +Open Scope monad_scope. +Existing Instances Monad_DS MonadLaws_DS MonadState_DS MonadZero_DS. + +(* The induction principle in the Coq standard library is much more awkward. *) +Lemma Z_nonneg_peano_ind (P : Z -> Prop) : + P 0%Z -> + (forall z, 0 <= z -> P z -> P (Z.succ z)) -> + (forall p, P (Z.neg p)) -> + forall z, P z. +Proof. + intros P0 Pind Pneg z. + destruct z. + - assumption. + - apply (Pos.peano_ind (fun p => P (Z.pos p))). + + change 1 with (Z.succ 0). + apply Pind. + * apply Z.le_refl. + * assumption. + + intros p' Pposp'. + rewrite Pos2Z.inj_succ. + apply Pind. + * apply Pos2Z.is_nonneg. + * assumption. + - apply Pneg. +Qed. + +Lemma Z_pos_peano_ind (P : Z -> Prop) : + (forall z, z <= 0 -> P z) -> + (forall z, 0 <= z -> P z -> P (Z.succ z)) -> + forall z, P z. +Proof. + intros Pnonpos Pind z. + destruct z. + - apply Pnonpos. omega. + - apply (Pos.peano_ind (fun p => P (Z.pos p))). + + change 1 with (Z.succ 0). + apply Pind; [|apply Pnonpos]; omega. + + intros p' Pposp'. + rewrite Pos2Z.inj_succ. + apply Pind; [apply Pos2Z.is_nonneg | assumption]. + - apply Pnonpos. apply Pos2Z.neg_is_nonpos. +Qed. + +(** Why we define a new iterator: Ziteri instead of Z.iter? + + Otherwise, when executing loop with runStateT, it will look like: + runStateT (fst (Z.iter ...(init_monad, start))) m. + We need to carry a number (ranges from start to bound) during iteration, + it can only be stored with the actual monad and form a pair. + This form is hard for applying MonadLaws. + So we let Ziteri itself to provide iteration value. +*) + +Fixpoint Niteri {A:Type} (f:Z->A->A) (n:nat) (x:A) {struct n} : A := + match n with + | O => x + | S n' => f (Z.of_nat n') (Niteri f n' x) + end. + +Lemma Niteri_swap: + forall (n:nat) (A:Type) (f:Z->A->A) (x:A), + Niteri (fun z => f (Z.succ z)) n (f 0 x) = f (Z.of_nat n) (Niteri f n x). +Proof. + induction n; intros. + - tauto. + - simpl. + rewrite IHn. + rewrite Zpos_P_of_succ_nat. reflexivity. +Qed. + +Lemma Niteri_succ: + forall n (A:Type) (f:Z->A->A) (x:A), + Niteri f (S n) x = f (Z.of_nat n) (Niteri f n x). +Proof. + reflexivity. +Qed. + +Lemma Niteri_add: + forall p q (A:Type) (f:Z->A->A) (x:A), + Niteri f (p+q) x = Niteri (fun z => f (z+(Z.of_nat q))) p (Niteri f q x). +Proof. + induction p; intros. + - reflexivity. + - (*now rewrite plus_Sn_m, !Niteri_succ, IHp. *) + rewrite plus_Sn_m. rewrite !Niteri_succ. rewrite IHp. + replace (Z.of_nat (p+q)) with (Z.of_nat p + Z.of_nat q). + reflexivity. + rewrite <- Nat2Z.inj_add. reflexivity. +Qed. + +Lemma Niteri_add1: + forall p (A:Type) (f:Z->A->A) (x:A), + Niteri f (S p) x = Niteri (fun z => f (Z.succ z)) p (Niteri f 1 x). +Proof. + intros. + replace (S p) with (p + 1)%nat by (rewrite Plus.plus_comm; reflexivity). + rewrite Niteri_add. reflexivity. +Qed. + +Definition Ziteri {A} (f:Z->A->A) (steps:Z) (x:A) := + Niteri f (Z.to_nat steps) x. +(* TODO: change all _ > _ or _ >= _ into _ < _ and _ <= _ *) +Lemma Ziteri_base': + forall z (A:Type) (f:Z->A->A) (x:A), + z < 0 -> + Ziteri f z x = x. +Proof. + intros z A f x lt0. + destruct z. + - omega. + - inversion lt0. + - unfold Ziteri. rewrite Z2Nat.inj_neg. reflexivity. +Qed. + +Lemma Ziteri_base: + forall z (A:Type) (f:Z->A->A) (x:A), + z <= 0 -> + Ziteri f z x = x. +Proof. + intros z A f x le0. + destruct z. + - reflexivity. + - destruct le0. reflexivity. + - unfold Ziteri. rewrite Z2Nat.inj_neg. reflexivity. +Qed. + +Lemma Ziteri_neg: + forall p (A:Type) (f:Z->A->A) (x:A), + Ziteri f (Zneg p) x = x. +Proof. + intros. + apply Ziteri_base. + specialize (Zlt_neg_0 p). + omega. +Qed. + +Lemma Ziteri_0: + forall (A:Type) (f:Z->A->A) (x:A), + Ziteri f 0 x = x. +Proof. + intros. + apply Ziteri_base. omega. +Qed. + +Lemma Ziteri_swap: + forall z (A:Type) (f:Z->A->A) (x:A), + 0 <= z -> + Ziteri (fun z => f (Z.succ z)) z (f 0 x) = f z (Ziteri f z x). +Proof. + intro z. + intros. unfold Ziteri. + rewrite (Niteri_swap (Z.to_nat z)). + rewrite Z2Nat.id. reflexivity. + assumption. +Qed. + +Lemma Ziteri_succ: + forall z (A:Type) (f:Z->A->A) (x:A), + 0 <= z -> + Ziteri f (Z.succ z) x = f z (Ziteri f z x). +Proof. + intros z A f x z_le0. + unfold Ziteri. + rewrite Z2Nat.inj_succ by assumption. + rewrite <- (Z2Nat.id z) at 2 by assumption. + apply Niteri_succ. +Qed. + +Lemma Ziteri_add1: + forall p (A:Type) (f:Z->A->A) (x:A), + 0 <= p -> + Ziteri f (Z.succ p) x = Ziteri (fun z => f (Z.succ z)) p (Ziteri f 1 x). +Proof. + intros. + unfold Ziteri. + rewrite Z2Nat.inj_succ by assumption. + apply Niteri_add1. +Qed. + + +Definition failure_passing {A} (h: DS A -> DS A) : Prop := + forall m x, + runStateT x m = mzero -> + runStateT (h x) m = mzero. + +Lemma Niteri_from_none {A} (f: Z -> DS A -> DS A) (fp: forall n, failure_passing (f n)) n: + forall x m, + runStateT x m = mzero -> + runStateT (Niteri f n x) m = mzero. +Proof. + induction n; intros x m opt_x. + - (* O *) + simpl. exact opt_x. + - (* S *) + simpl. + apply fp. + apply IHn. exact opt_x. +Qed. + +Lemma Niteri_cut_0 {A} {f: Z -> DS A -> DS A} (fp: forall n, failure_passing (f n)) {n x m v s}: + runStateT (Niteri f n x) m = ret (v, s) -> + exists v' s', runStateT x m = ret (v', s'). +Proof. + intros opt_n. + destruct (runStateT x m) eqn:opt_x. + - (* run x m = Some p *) + destruct p as [v' s']. exists v', s'. reflexivity. + - (* run x m = None *) + apply (Niteri_from_none f fp n) in opt_x. + rewrite opt_x in opt_n. + discriminate. +Qed. + +Lemma Niteri_Ziteri {A} (f: Z -> DS A -> DS A) {P}: + forall x n, + P (Niteri f (Z.to_nat n) x) -> + P (Ziteri f n x). +Proof. + intros x n Pn. + unfold Ziteri. + exact Pn. +Qed. + +Lemma Ziteri_from_none {A} (f: Z -> DS A -> DS A) (fp: forall n, failure_passing (f n)) n: + forall x m, + runStateT x m = mzero -> + runStateT (Ziteri f n x) m = mzero. +Proof. + intros x m opt_x. + apply (Niteri_from_none f fp (Z.to_nat n)) in opt_x. + unfold Ziteri. + exact opt_x. +Qed. + +Lemma Ziteri_cut_0 {A} {f: Z -> DS A -> DS A} (fp: forall n, failure_passing (f n)) {n x m v s}: + runStateT (Ziteri f n x) m = ret (v, s) -> + exists v' s', runStateT x m = ret (v', s'). +Proof. + unfold Ziteri. + intros opt_n. + apply Niteri_Ziteri in opt_n. + apply (Niteri_cut_0 fp) in opt_n. + exact opt_n. +Qed. + +Definition oZiteri_wrapper {A} (f: Z -> A -> DS A) n x : DS A := + r <- x;; f n r. + +Lemma oZiteri_wrapper_unfolder {A} (f: Z -> A -> DS A) n x: + oZiteri_wrapper f n x = (r <- x;; f n r). + reflexivity. +Qed. + +Definition oZiteri {A} (f: Z -> A -> DS A) n x : DS A := + Ziteri (oZiteri_wrapper f) n (ret x). + +Definition oNiteri {A} (f: Z -> A -> DS A) n x : DS A := + Niteri (oZiteri_wrapper f) n (ret x). + +Lemma oNiteri_unfolder {A} (f: Z -> A -> DS A) n x: + oNiteri f n x = Niteri (oZiteri_wrapper f) n (ret x). + reflexivity. +Qed. + +Lemma oZiteri_unfolder {A} (f: Z -> A -> DS A) n x: + oZiteri f n x = Ziteri (oZiteri_wrapper f) n (ret x). + reflexivity. +Qed. + +Lemma oZiteri_oNiteri {A} (f: Z -> A -> DS A) n x: + oZiteri f n x = oNiteri f (Z.to_nat n) x. +Proof. + reflexivity. +Qed. + +Definition first_map (check: Z -> DS (tp_ft int_bool_pair)) + (start bound : Z) := + oZiteri + ( fun i l => m <- get;; h <- check (start + i);; put m) + (bound - start) + tt. + +Definition fold_step_opt {htp tmp id_it id_recur}(se : spec_env_t tmp) (start:Z): + let tmp' := AList.set id_recur htp (AList.set id_it int_Z32_pair tmp) in + (SpecTree.t tmp' -> DS (tp_ft htp)) -> + Z -> + tp_ft htp -> + DS (tp_ft htp) := + fun c i r => + c (SpecTree.set id_recur htp r + (SpecTree.set id_it int_Z32_pair (start+i) se)). + +Lemma oNiteri_link_head {A} {f: Z -> A -> DS A} {n x m v s}: + runStateT (oNiteri f 1 x) m = ret (v, s) -> + runStateT (oNiteri (fun z => f (Z.succ z)) n v) s = + runStateT (oNiteri f (S n) x) m. +Proof. + revert x m v s. + induction n; intros x m v s opt_1. + - (* O *) + unfold oNiteri at 1. simpl. rewrite opt_1. reflexivity. + - (* S n *) + specialize (IHn _ _ _ _ opt_1). + unfold oNiteri at 1. + simpl. + unfold oZiteri_wrapper at 1. + rewrite <- oNiteri_unfolder. + bind_rewrite_osT IHn. + rewrite <- Nat2Z.inj_succ. + reflexivity. +Qed. + +Lemma oZiteri_link_head {A} {f: Z -> A -> DS A} {n x m v s}: + 0 <= n -> + runStateT (oZiteri f 1 x) m = ret (v, s) -> + runStateT (oZiteri (fun z => f (Z.succ z)) n v) s = + runStateT (oZiteri f (Z.succ n) x) m. +Proof. + intros n_range. + repeat rewrite oZiteri_oNiteri. + replace (Z.to_nat 1) with 1%nat by reflexivity. + replace (Z.to_nat (Z.succ n)) with (S (Z.to_nat n)). + exact oNiteri_link_head. + symmetry. apply Z2Nat.inj_succ. assumption. +Qed. + +Lemma failure_passing_oZiteri_wrapper {A} (f: Z -> A -> DS A): + forall n, failure_passing ((oZiteri_wrapper f) n). +Proof. + intros n. + unfold oZiteri_wrapper, failure_passing. + intros m x opt_x. + apply bind_rewrite_osT_mzero. exact opt_x. +Qed. + +Lemma oZiteri_0 {A} {f: Z -> A -> DS A} {x:A}: + oZiteri f 0 x = ret x. +Proof. + unfold oZiteri. apply Ziteri_0. +Qed. + +Lemma oZiteri_base {A} {f: Z -> A -> DS A} {x:A} n: + n <= 0 -> + oZiteri f n x = ret x. +Proof. + intros n_range. + unfold oZiteri. apply Ziteri_base. assumption. +Qed. + +Lemma oZiteri_cut_head {A} {f: Z -> A -> DS A} {n x m v s}: + 0 <= n -> + runStateT (oZiteri f (Z.succ n) x) m = ret (v, s) -> + exists v' s', runStateT (oZiteri f 1 x) m = ret (v', s') /\ + runStateT (oZiteri (fun z => f (Z.succ z)) n v') s' = ret (v, s). +Proof. + intros n_range opt_succ_n. + unfold oZiteri in opt_succ_n |-*. + rewrite Ziteri_add1 in opt_succ_n by assumption. + assert (fp: forall n0 : Z, failure_passing (oZiteri_wrapper f (Z.succ n0))). + { apply failure_passing_oZiteri_wrapper. } + destruct (Ziteri_cut_0 fp opt_succ_n) as (v' & s' & opt_1). + exists v', s'. + split; [assumption|]. + rewrite <- oZiteri_unfolder. + rewrite <- Ziteri_add1 in opt_succ_n by assumption. + rewrite <- oZiteri_unfolder in opt_succ_n, opt_1. + erewrite oZiteri_link_head. + exact opt_succ_n. + omega. + exact opt_1. +Qed. + +Lemma Ziteri_1 {A} {f: Z -> A -> A} {x:A}: + Ziteri f 1 x = f 0 x. +Proof. + reflexivity. +Qed. + +Lemma oZiteri_1 {A} {f: Z -> A -> DS A} {x m}: + runStateT (oZiteri f 1 x) m = runStateT (f 0 x) m. +Proof. + unfold oZiteri. + rewrite Ziteri_1. + unfold oZiteri_wrapper. + destruct (runStateT (f 0 x) m) eqn:opt_0. + - destruct p as [v s]. + erewrite bind_osT. + reflexivity. + apply ret_osT. split; reflexivity. + exact opt_0. + - rewrite bind_of_return. exact opt_0. + typeclasses eauto. +Qed. + +Lemma oZiteri_neg {A} {f: Z -> A -> DS A} {x:A} {p}: + oZiteri f (Z.neg p) x = ret x. +Proof. + reflexivity. +Qed. + +Lemma oZiteri_cut_tail {A} {f: Z -> A -> DS A} {n x m v s}: + 0 <= n -> + runStateT (oZiteri f (Z.succ n) x) m = ret (v, s) -> + exists v' s', runStateT (oZiteri f n x) m = ret (v', s') /\ + runStateT (oZiteri (fun z => f (n + z)) 1 v') s' = ret (v, s). +Proof. + intros n_range opt_succ_n. + unfold oZiteri in opt_succ_n |-*. + rewrite Ziteri_succ in opt_succ_n by assumption. + unfold oZiteri_wrapper at 1 in opt_succ_n. + split_run opt_succ_n as (v' & s' & opt_n & opt_rest). + exists v', s'. + split. exact opt_n. + rewrite Ziteri_1. + unfold oZiteri_wrapper. + replace (n + 0) with n by apply Zplus_0_r_reverse. + erewrite bind_osT. + - reflexivity. + - apply ret_osT. split; reflexivity. + - exact opt_rest. +Qed. +Lemma oZiteri_mid {A} {f: Z -> A -> DS A} n0 n1 {x m v s}: + 0 <= n0 -> + 0 <= n1 -> + runStateT (oZiteri f (n0+n1) x) m = ret (v, s) -> + exists v' s', runStateT (oZiteri f n0 x) m = ret (v', s') /\ + runStateT (oZiteri (fun z => f (n0 + z)) n1 v') s' = ret (v, s). +Proof. + revert n1 n0 x m v s. + refine (Z_nonneg_peano_ind _ _ _ _). + - (* n1 = 0 *) + intros n0 x m v s n0_range n1_range opt. (* 0 = n1 *) + replace (n0 + 0) with n0 in opt by apply Zplus_0_r_reverse. + exists v, s. + split. assumption. + reflexivity. + - (* n1 > 0 *) + intros z z_range Hind n0 x m v s n0_range n1_range opt. (* Z.succ z = n1 *) + replace (n0 + Z.succ z) with (Z.succ (n0 + z)) in opt by omega. + apply oZiteri_cut_tail in opt. 2:omega. + destruct opt as (v1 & s1 & opt_n0_z & opt_rest). + destruct (Hind _ _ _ _ _ n0_range z_range opt_n0_z) as (v' & s' & opt_n0 & opt_z). + exists v', s'. + split; [assumption|]. + (* NOTE: unfold oZiteri in opt_rest *) + (* unfold oZiteri in opt_rest. *) + (* rewrite Ziteri_1 in opt_rest. *) + unfold oZiteri. + rewrite Ziteri_succ by assumption. + rewrite <- oZiteri_unfolder. + unfold oZiteri_wrapper. + erewrite bind_osT. + + reflexivity. + + exact opt_z. + + rewrite oZiteri_1 in opt_rest. + replace (n0 + z + 0) with (n0 + z) in opt_rest by omega. + exact opt_rest. + - (* n1 < 0 *) + intros p n0 x m v s n0_range n1_range opt. (* Z.neg p = n1 *) + pose (Pos2Z.neg_is_neg p) as n1_neg. + omega. +Qed. + +Lemma oZiteri_pure {A} {f: Z -> A -> DS A}: + forall n m, + (forall z, 0 <= z < n -> forall r v s, runStateT (f z r) m = ret (v, s) -> m = s) -> + forall x v s, runStateT (oZiteri f n x) m = ret (v, s) -> + m = s. +Proof. + refine (Z_nonneg_peano_ind _ _ _ _). + - (* n = 0 *) + intros m pure_step x v s opt. + rewrite oZiteri_0 in opt. + apply ret_osT in opt. tauto. + - (* n > 0 *) + intros z z_range Hind m pure_step x v s opt. + apply oZiteri_cut_tail in opt. 2:assumption. + destruct opt as (v' & s' & opt_z & opt_last). + assert (m = s'). + { eapply Hind. 2:exact opt_z. + intros z0 z0_range. apply pure_step. omega. } + subst s'. + rewrite oZiteri_1 in opt_last. + replace (z + 0) with z in opt_last by omega. + eapply (pure_step z). omega. + exact opt_last. + - (* n < 0 *) + intros p m pure_step x v s opt. + rewrite oZiteri_neg in opt. + apply ret_osT in opt. tauto. +Qed. + +(* connect multiple individual steps into n steps *) +Lemma oZiteri_connect {A} {f: Z -> A -> DS A}: + forall n x m, + (forall z, 0 <= z < n -> + forall v' s', runStateT (oZiteri f z x) m = ret (v', s') -> + exists v s, runStateT (f z v') s' = ret (v, s)) -> + exists v s, runStateT (oZiteri f n x) m = ret (v, s). +Proof. + intros n. revert n f. + refine (Z_nonneg_peano_ind _ _ _ _). + - (* n = 0 *) + intros f x m opt_step. + exists x, m. + reflexivity. + - (* n > 0 *) + intros z z_range Hind f x m opt_step. + assert (opt_1: exists v1 s1, runStateT (f 0 x) m = ret (v1, s1)). + { eapply opt_step. omega. + rewrite oZiteri_0. reflexivity. } + destruct opt_1 as (v1 & s1 & opt_1). + assert (opt_rest: exists v s, runStateT (oZiteri (fun z => f (Z.succ z)) z v1) s1 = ret (v, s)). + { apply Hind. + intros z0 z0_range v' s' opt_z0. + apply opt_step. omega. + erewrite <- oZiteri_link_head. + - exact opt_z0. + - apply z0_range. + - rewrite oZiteri_1. exact opt_1. } + destruct opt_rest as (v & s & opt_rest). + exists v, s. + erewrite <- oZiteri_link_head. + + exact opt_rest. + + assumption. + + rewrite oZiteri_1. exact opt_1. + - (* n < 0 *) + intros p f x m opt_step. + exists x, m. + reflexivity. +Qed. + +(* TODO: This proof is not quite elegant, should have been proven by oZiteri_pure and oZiteri_connect. *) +Lemma oZiteri_pure_connect {A} {f: Z -> A -> DS A}: + forall n x m, + (forall z, 0 <= z < n -> + forall v', runStateT (oZiteri f z x) m = ret (v', m) -> + exists v, runStateT (f z v') m = ret (v, m)) -> + exists v, runStateT (oZiteri f n x) m = ret (v, m). +Proof. + intros n. revert n f. + refine (Z_nonneg_peano_ind _ _ _ _). + - (* n = 0 *) + intros f x m opt_step. + exists x. + reflexivity. + - (* n > 0 *) + intros z z_range Hind f x m opt_step. + assert (opt_1: exists v1, runStateT (f 0 x) m = ret (v1, m)). + { eapply opt_step. omega. + rewrite oZiteri_0. reflexivity. } + destruct opt_1 as [v1 opt_1]. + assert (opt_rest: exists v, runStateT (oZiteri (fun z => f (Z.succ z)) z v1) m = ret (v, m)). + { apply Hind. + intros z0 z0_range v' opt_z0. + apply opt_step. omega. + erewrite <- oZiteri_link_head. + - exact opt_z0. + - apply z0_range. + - rewrite oZiteri_1. exact opt_1. } + destruct opt_rest as [v opt_rest]. + exists v. + erewrite <- oZiteri_link_head. + + exact opt_rest. + + assumption. + + rewrite oZiteri_1. exact opt_1. + - (* n < 0 *) + intros p f x m opt_step. + exists x. + reflexivity. +Qed. + +Definition for_step_opt {tmp} {id_it : positive} (se : spec_env_t tmp) (start:Z) : + let tmp' := AList.set id_it int_Z32_pair tmp in + (spec_env_t tmp' -> DS unit) -> + Z -> + DS (tp_ft void_unit_pair) -> + DS (tp_ft void_unit_pair) := + fun c i ma => + ma;; c (SpecTree.set id_it int_Z32_pair (start+i) se). + +Lemma for_from_none {tmp} id_it (se: spec_env_t tmp): + forall p f s0 start c, + runStateT c s0 = mzero -> + runStateT (Niteri (@for_step_opt _ id_it se start f) + p + c) s0 = mzero. +Proof. + induction p; intros. + - tauto. + - simpl. + unfold for_step_opt at 1. + apply bind_rewrite_osT_mzero. + apply IHp. assumption. +Qed. + +Lemma for_from_none' {tmp} id_it (se: spec_env_t tmp): + forall p f s0 start init, + runStateT init s0 = mzero -> + runStateT (Niteri (fun z => @for_step_opt _ id_it se start f (Z.succ z)) + p + init) s0 = mzero. +Proof. + induction p; intros. + - tauto. + - simpl. + unfold for_step_opt at 1. + apply bind_rewrite_osT_mzero. + apply IHp. assumption. +Qed. + +Lemma for_n_plus_1__runStateT: + forall tmp id_it (se:spec_env_t tmp) + (f: SpecTree.ilist (AList.set id_it int_Z32_pair tmp) -> stateT GetHighData option (tp_ft void_unit_pair)) d start mh m', + runStateT (f (SpecTree.set id_it int_Z32_pair start se)) mh = ret (tt, m') -> + runStateT + (Niteri + (fun z : Z => + @for_step_opt _ id_it se start f (Z.succ z)) d + (Niteri (for_step_opt se start f) 1 (ret tt))) + mh = + runStateT + (Niteri (for_step_opt se (start+1) f) d (ret tt)) + m'. +Proof. + induction d; intros. + - simpl. + unfold for_step_opt at 1. + rewrite Z.add_0_r. apply H. + - simpl. simpl in IHd. + unfold for_step_opt at 1. + unfold for_step_opt at 3. + specialize (IHd _ _ _ H). + bind_rewrite_osT IHd. + replace (start + Z.succ (Z.of_nat d)) with (start + 1 + Z.of_nat d) by omega. + reflexivity. +Qed. + +Lemma for_n_plus_1__execStateT: + forall tmp id_it (se:spec_env_t tmp) (f: SpecTree.ilist (AList.set id_it int_Z32_pair tmp) -> stateT GetHighData option (tp_ft void_unit_pair)) d start mh m', + execStateT (f (SpecTree.set id_it int_Z32_pair start se)) mh = ret m' -> + execStateT + (Ziteri + (fun z : Z => + @for_step_opt _ id_it se start f (Z.succ z)) d + (Ziteri (for_step_opt se start f) 1 (ret tt))) + mh = + execStateT + (Ziteri (for_step_opt se (start+1) f) d (ret tt)) + m'. +Proof. + unfold Ziteri. + intros. + Transparent execStateT. + unfold execStateT. + Opaque execStateT. + rewrite (for_n_plus_1__runStateT _ _ _ _ _ start mh m'). + reflexivity. + exec2run H as v. + destruct v; auto. +Qed. + +Lemma for_cut_opt' {tmp id_it} (se : spec_env_t tmp): + let tmp' := AList.set id_it int_Z32_pair tmp in + forall f s0 start steps s, + 0 <= steps -> + runStateT (Ziteri (@for_step_opt _ id_it se start f) + (Z.succ steps) + (ret tt)) s0 = ret (tt, s) -> + exists s', + runStateT (f (SpecTree.set id_it int_Z32_pair start se)) s0 = ret (tt, s') + /\ runStateT (Ziteri (@for_step_opt _ id_it se (start+1) f) + steps + (ret tt)) s' = ret (tt, s). +Proof. + intros until s. intros steps_nonneg Hiter. + destruct (runStateT (f (SpecTree.set id_it int_Z32_pair start se)) s0) eqn:Hrun. + - destruct p as [v' s']; exists s'; destruct v'; split; [reflexivity|]. + rewrite Ziteri_add1 in Hiter by assumption. + unfold Ziteri in Hiter. + rewrite (for_n_plus_1__runStateT _ _ _ _ _ _ _ _ Hrun) in Hiter. + apply Hiter. + + - simpl in Hiter. + unfold Ziteri in Hiter. + rewrite Z2Nat.inj_succ in Hiter by assumption. + rewrite Niteri_add1 in Hiter. + rewrite for_from_none' in Hiter. discriminate. + simpl. unfold for_step_opt at 1. + rewrite bind_of_return by apply MonadLaws_DS. + rewrite Z.add_0_r. assumption. +Qed. + +Lemma oZiteri_fail_at_1 {A} {f: Z -> A -> DS A} {n x m}: + 0 <= n -> + runStateT (oZiteri f 1 x) m = mzero -> + runStateT (oZiteri f (Z.succ n) x) m = mzero. +Proof. + intros n_range opt_1. + rewrite oZiteri_1 in opt_1. + unfold oZiteri. + rewrite Ziteri_add1 by assumption. + apply Ziteri_from_none. + { (* failure_passing *) + unfold failure_passing, oZiteri_wrapper. + intros n0 m0 x0 opt_x0. + apply bind_rewrite_osT_mzero. assumption. } + rewrite Ziteri_1. unfold oZiteri_wrapper. + erewrite bind_rewrite_osT_ret. exact opt_1. reflexivity. +Qed. + +Lemma oZiteri_replace {A} {f: Z -> A -> DS A} {g: Z -> A -> DS A}: + (forall z r s, runStateT (f z r) s = runStateT (g z r) s) -> + forall n x m, runStateT (oZiteri f n x) m = runStateT (oZiteri g n x) m. +Proof. + intros Heq n. revert n f g Heq. + refine (Z_nonneg_peano_ind _ _ _ _). + - (* n = 0 *) + intros f g Heq x m. reflexivity. + - (* n > 0 *) + intros z z_range Hind f g Heq x m. + destruct (runStateT (oZiteri f 1 x) m) eqn:opt_1. + + destruct p as [v s]. + erewrite <- (oZiteri_link_head z_range opt_1). + rewrite oZiteri_1 in opt_1. + rewrite Heq in opt_1. + rewrite <- oZiteri_1 in opt_1. + erewrite <- (oZiteri_link_head z_range opt_1). + eapply Hind. + intros z0 r s0. + apply Heq. + + rewrite oZiteri_fail_at_1 by assumption. + symmetry. + apply oZiteri_fail_at_1. assumption. + rewrite oZiteri_1 in opt_1. rewrite Heq in opt_1. + exact opt_1. + - (* n < 0 *) + intros p f g Heq x m. + do 2 rewrite oZiteri_neg. + reflexivity. +Qed. + +Lemma fold_n_plus_1__runStateT {htp tmp id_it id_recur} (se:spec_env_t tmp) f mh start init: + let tmp''' := AList.set id_recur htp (AList.set id_it int_Z32_pair tmp) in + let se_recur := SpecTree.set id_recur htp init (SpecTree.set id_it int_Z32_pair start se) in + forall v' s', + runStateT (f se_recur) mh = ret (v', s') -> + forall n, + 0 <= n -> + runStateT + (oZiteri + (fold_step_opt se start f) (Z.succ n) init) mh = + runStateT + (oZiteri (fold_step_opt se (Z.succ start) f) n v') s'. +Proof. + intros until s'. intros opt_f n n_range. + unfold fold_step_opt. + erewrite <- oZiteri_link_head. + - apply oZiteri_replace. + intros z r s. + replace (start + Z.succ z) with (Z.succ start + z) by omega. reflexivity. + - assumption. + - rewrite oZiteri_1. replace (start + 0) with start by apply Zplus_0_r_reverse. apply opt_f. +Qed. + +Lemma fold_cut_opt' {htp tmp id_it id_recur} (se : spec_env_t tmp) f mh start init: + let tmp''' := AList.set id_recur htp (AList.set id_it int_Z32_pair tmp) in + let se_recur := SpecTree.set id_recur htp init (SpecTree.set id_it int_Z32_pair start se) in + forall steps v s, + 0 <= steps -> + runStateT (oZiteri (@fold_step_opt _ _ id_it id_recur se start f) + (Z.succ steps) + init) mh = ret (v, s) -> + exists v' s', + runStateT (f se_recur) mh = ret (v', s') /\ + runStateT (oZiteri (@fold_step_opt _ _ id_it id_recur se (Z.succ start) f) + steps + v') s' = ret (v, s). +Proof. + intros until s. intros steps_nonneg opt_succ_steps. + apply oZiteri_cut_head in opt_succ_steps. 2:assumption. + destruct opt_succ_steps as (v' & s' & opt_1 & opt_steps). + exists v', s'. + rewrite oZiteri_1 in opt_1. unfold fold_step_opt in opt_1. + replace (start + 0) with start in opt_1 by apply Zplus_0_r_reverse. + split. assumption. + erewrite oZiteri_replace. exact opt_steps. + intros z r s0. + unfold fold_step_opt. + rewrite Z.add_succ_comm. reflexivity. +Qed. + + +Lemma for_cut_opt {tmp id_it} (se : spec_env_t tmp): + let tmp' := AList.set id_it int_Z32_pair tmp in + forall f s0 start steps s, + 0 <= steps -> + execStateT (Ziteri (@for_step_opt _ id_it se start f) + (Z.succ steps) + (ret tt)) s0 = ret s -> + exists s', + execStateT (f (SpecTree.set id_it int_Z32_pair start se)) s0 = ret s' + /\ execStateT (Ziteri (@for_step_opt _ id_it se (start+1) f) + steps + (ret tt)) s' = ret s. +Proof. + intros ? ? ? ? ? ? steps_nonneg Hiter. + simpl in *. + exec2run Hiter as []. + apply for_cut_opt' in Hiter; [|assumption]. + destruct Hiter as (s' & Hiter_1 & Hiter_n). + apply runStateT_execStateT in Hiter_1. + apply runStateT_execStateT in Hiter_n. + exists s'. split. + - exact Hiter_1. + - exact Hiter_n. +Qed. + +Fixpoint Nfind (check:Z -> DS bool) (n:nat) {struct n} : DS (option Z) := + match n with + | O => ret None + | S n' => res <- check (Z.of_nat n);; + if res + then ret (Some (Z.of_nat n)) + else Nfind check n' + end. + +Definition Zfind (check:Z -> DS bool) (z:Z) := + Nfind check (Z.to_nat z). + +Lemma Zfind_succ (check:Z -> DS bool) (z:Z): + 0 <= z -> + Zfind check (Z.succ z) = + (res <- check (Z.succ z);; + if res + then ret (Some (Z.succ z)) + else Zfind check z). +Proof. + intro z_nonneg. + unfold Zfind. + rewrite Z2Nat.inj_succ by assumption. + replace (Z.succ z) with (Z.of_nat (S (Z.to_nat z))) by + (rewrite <- (Z2Nat.id z) at 2 by assumption; + apply Nat2Z.inj_succ). + reflexivity. +Qed. + +(* Transparent ret. *) +Lemma Zfind_pure (check:Z -> DS bool): + (forall z m v s, runStateT (check z) m = ret (v, s) -> m = s) -> + forall z m v s, runStateT (Zfind check z) m = ret (v, s) -> + m = s. +Proof. + intros Hstep. + refine (Z_nonneg_peano_ind _ _ _ _). + - (* z = 0 *) + intros ? ? ? Hrun. + unfold Zfind in Hrun. simpl in Hrun. + apply ret_osT in Hrun. apply Hrun. + - (* z > 0 *) + intros ? Hz Hind ? ? ? Hrun. + rewrite Zfind_succ in Hrun by omega. + dest_opt Hrun as Hrun' [v' s']. + specialize (Hstep (Z.succ z) m v' s'). + apply Hstep in Hrun; subst s'; clear Hstep. + destruct v'. + + (* check = true *) + apply ret_osT in Hrun'. apply Hrun'. + + (* check = false *) + eapply Hind. exact Hrun'. + - (* z < 0 *) + intros ? ? ? ? Hrun. + unfold Zfind in Hrun. simpl in Hrun. + apply ret_osT in Hrun. apply Hrun. +Qed. + +Definition first_spec {A} + (check: Z -> DS (tp_ft int_bool_pair)) (f: Z -> A) + (g: Z -> A) start bound (m:GetHighData) := + let P n := is_true + (match evalStateT (check n) m with + | Some b => b + | None => false + end) + in + let Pdec n : {P n} + {~ P n} := bool_dec _ true in + match LangDef.min_ex P start bound Pdec with + | inleft (exist n _) => f n + | inright _ => g bound + end. + +Ltac dest_first H := + try unfold first_spec in H; + match type of H with + | runStateT match ?X with _ => _ end _ = _ => + let n := fresh "n" in + let n_range := fresh "n_range" in + let Pn := fresh "Pn" in + let n_least := fresh "n_least" in + let no_n := fresh "no_n" in + destruct X as [(n & n_range & Pn & n_least) | no_n] + end. + +Ltac dest_first_eqn H := + try unfold first_spec in H; + match type of H with + | runStateT match ?X with _ => _ end _ = _ => + let n := fresh "n" in + let n_range := fresh "n_range" in + let Pn := fresh "Pn" in + let n_least := fresh "n_least" in + let no_n := fresh "no_n" in + let first_eq := fresh "first_eq" in + destruct X as [(n & n_range & Pn & n_least) | no_n] eqn:first_eq + end. + +Definition ext_call_me (ext_contract : int256) := {| + me_address := ext_contract; + me_origin := me_origin me; + me_caller := me_address me; + me_callvalue := me_callvalue me; (* FIXME: the callvalue modeling is wrong *) + me_coinbase := me_coinbase me; + me_timestamp := me_timestamp me; + me_number := me_number me; + me_chainid := me_chainid me; + me_selfbalance := me_selfbalance me; + me_balance := me_balance me; + me_blockhash := me_blockhash me; + me_transfer := me_transfer me; + me_callmethod := me_callmethod me; + me_log := me_log me; +|}. + +Fixpoint synth_stmt_spec_opt {returns}(c : cmd_constr returns) dest tmp : + synth_stmt_wellformed c dest tmp -> spec_env_t tmp -> DS (tp_ft returns) := + match c with + | CCskip => fun wf se => ret tt + | CClet _r tp _ id c1 c2 => + let htp := mk_hyper_type_pair tp in + (fun (wf : let tmp' := AList.set id (mk_hyper_type_pair tp) tmp in + ((tmp ! id ~~~ None) * dest_not_exist dest tmp') * + (synth_stmt_wellformed c1 id tmp * synth_stmt_wellformed c2 dest tmp')) + se => + spec1 <- (synth_stmt_spec_opt c1 id tmp (cadr wf) se);; + let se' := SpecTree.set id htp spec1 se in + (synth_stmt_spec_opt c2 dest _ (cddr wf) se')) + | CCload tp _ e => fun wf se => + gets (synth_lexpr_spec me tmp e wf se).(ltype_get) + | CCstore _ _ el er => fun wf se => + let f := synth_expr_spec me tmp er (cdr wf) se in + modify ((synth_lexpr_spec me tmp el (car wf) se).(ltype_set) f) + | CCsequence _ c1 c2 => + fun (wf : synth_stmt_wellformed c1 dest tmp * synth_stmt_wellformed c2 dest tmp) + se => + (synth_stmt_spec_opt c1 dest tmp (car wf) se) ;; + (synth_stmt_spec_opt c2 dest tmp (cdr wf) se) + | CCifthenelse _ e c_true c_false => + fun (wf : synth_expr_wellformed tmp e * + (synth_stmt_wellformed c_true dest tmp * synth_stmt_wellformed c_false dest tmp)) + se => + if synth_expr_spec me tmp e (car wf) se + then synth_stmt_spec_opt c_true dest tmp (cadr wf) se + else synth_stmt_spec_opt c_false dest tmp (cddr wf) se + | CCfor id_it id_end e1 e2 c3 => + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + dest_not_exist dest tmp'')) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + synth_stmt_wellformed c3 dest tmp''))) + se => + let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + Ziteri (for_step_opt initial_se start + (synth_stmt_spec_opt c3 dest _ (cdddr wf))) + (bound - start) + (ret tt) + | CCfirst _ id_it id_end id_dest e1 e2 c3 c4 c5 => + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + ((tmp'' ! id_dest ~~~ None) * (dest_not_exist dest tmp'' * + ((synth_stmt_pure c3 ~~~ true)))))) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + (synth_stmt_wellformed c3 id_dest tmp'' * + (synth_stmt_wellformed c4 dest tmp'' * + synth_stmt_wellformed c5 dest tmp''))))) + se => + let start := (synth_expr_spec me tmp e1 (cadr wf) se) in + let bound := (synth_expr_spec me tmp e2 (caddr wf) se) in + m <- get;; + first_map + (fun n => + synth_stmt_spec_opt c3 id_dest _ (cadddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se))) + start bound;; + first_spec + (fun n => + synth_stmt_spec_opt c3 id_dest _ (cadddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se))) + (fun n => + synth_stmt_spec_opt c4 dest _ (caddddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se))) + (fun n => + synth_stmt_spec_opt c5 dest _ (cdddddr wf) + (SpecTree.set id_it int_Z32_pair bound (* make more sense to swap n and bound here, + but I did not change it to consist with legacy code. *) + (SpecTree.set id_end int_Z32_pair n se))) + start + bound + m + | CCfold tp _ id_it id_end id_recur id_dest e1 e2 e3 c4 => + let htp := mk_hyper_type_pair tp in + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + let tmp''' := AList.set id_recur (mk_hyper_type_pair tp) tmp'' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + ((tmp'' ! id_recur ~~~ None) * ((tmp''' ! id_dest ~~~ None) * + (dest_not_exist dest tmp''' * (synth_stmt_pure c4 ~~~ true)))))) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + (synth_expr_wellformed tmp e3 * synth_stmt_wellformed c4 id_dest tmp''')))) + se => + let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + let init := synth_expr_spec me tmp e3 (cadddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + oZiteri (fold_step_opt initial_se start + (synth_stmt_spec_opt c4 id_dest _ (cddddr wf))) + (bound - start) + init + | CCcall argt ret prim args => + fun wf se => + let argv := map2_synth_expr_spec args se wf in + prim.(PRIMsem_opt) argv me + (* | CCcall_ext argt ret addr prim args => + fun wf se => + let argv := map2_synth_expr_spec args se wf in + prim.(PRIMsem_opt) argv (ext_call_me addr) *) + | CCyield tp _ e => + fun wf se => + ret (synth_expr_spec me tmp e wf se) + | CCconstr _ _ _ _ el flds constr => + fun wf se => + let l := synth_lexpr_spec me tmp el (car wf) se in + modify (l.(ltype_set) + (apply_curried_func + constr + (map2_synth_expr_spec flds se (cdr wf)))) +(* m <- get;; + put (l.(ltype_set) + (apply_curried_func + constr + (map2_synth_expr_spec flds se (cdr wf))) m) *) + | CCassert c => + fun (wf: (synth_stmt_pure c ~~~ true) * synth_stmt_wellformed c dest tmp) se => + v <- synth_stmt_spec_opt c dest tmp (cdr wf) se;; + guard v + | CCdeny c => + fun (wf: (synth_stmt_pure c ~~~ true) * synth_stmt_wellformed c dest tmp) se => + v <- synth_stmt_spec_opt c dest tmp (cdr wf) se;; + guard (negb v) + (* | CCghost c => + synth_stmt_spec_opt c dest tmp *) + | CCpanic _ _ => fun _ _ => mzero + | CCrespec _ tmp' c spec => + fun (wf : ((tmp ~~~ tmp') * (synth_stmt_pure c ~~~ true)) * + synth_stmt_wellformed c dest tmp) se => + m <- get;; + v <- spec (iso_f (caar wf) se);; + put m;; + ret v + | CCrespec_opt _ tmp' c spec => + fun (wf : (tmp ~~~ tmp') * synth_stmt_wellformed c dest tmp) se => + spec me (iso_f (car wf) se) + (* | _ => fun wf se => mzero *) +end %alist. + +Definition synth_stmt_spec_ret {returns} (c : cmd_constr returns) dest tmp: + synth_stmt_wellformed c dest tmp -> spec_env_t tmp -> GetHighData -> tp_ft returns := + fun wf se m => + match evalStateT (synth_stmt_spec_opt c dest tmp wf se) m with + | Some v => v + | None => ht_default + end. + + +Lemma fold_pure htp id_it id_recur tmp + (c:SpecTree.t (AList.set id_recur htp + (AList.set id_it int_Z32_pair tmp)) -> DS (tp_ft htp)) + (init:tp_ft htp) steps start: + (forall m v s se, runStateT (c se) m = ret (v, s) -> m = s) -> + forall m v s se, runStateT (oZiteri (@fold_step_opt htp tmp id_it id_recur se start c) steps init) m = ret (v, s) -> + m = s. +Proof. + intros pure_step m v s se opt_steps. + eapply oZiteri_pure. + 2:exact opt_steps. + intros z z_range r v0 s0 opt_z. + unfold fold_step_opt in opt_z. + apply pure_step in opt_z. assumption. +Qed. + +Lemma get_put_pure {A} {s v d} {f: DS A}: + runStateT (m <- get;; f ;; put m) d = ret (v, s) -> + d = s. +Proof. + intros opt. + dest_opt opt as opt1 [v1 s1]. + apply get_osT in opt. destruct opt; subst v1 s1. + dest_opt opt1 as opt2 [v2 s2]. + apply put_osT in opt2. destruct opt2; subst v s. + reflexivity. +Qed. + +Lemma first_map_pure' (check: Z -> DS bool) start steps: + (* (forall i m v s, 0 <= i < steps -> (runStateT (check (start+i)) m = ret (v, s) -> m = s)) -> *) + forall m l s, runStateT (first_map check start (start + steps)) m = ret (l, s) -> + m = s. +Proof. + unfold first_map. replace (start + steps - start) with steps by omega. + intros (* step_pure *) m l s opt. + eapply oZiteri_pure. 2:exact opt. + intros z z_range r v s0 step_pure'. + simpl in step_pure'. + apply get_put_pure in step_pure'. assumption. +(* split_run step_pure' as (v' & s' & opt_check & opt_ret). + rewrite ret_osT in opt_ret. + destruct opt_ret; subst v s0. + eapply (step_pure z). assumption. + exact opt_check. *) +Qed. + +Lemma first_map_pure (check: Z -> DS bool) start bound: + (* (forall i m v s, 0 <= i < (bound-start) -> runStateT (check (start+i)) m = ret (v, s) -> m = s) -> *) + forall m l s, runStateT (first_map check start bound) m = ret (l, s) -> + m = s. +Proof. + intros Hstep ? ? Hrun. + eapply first_map_pure'. + replace bound with (start + (bound - start)) in Hrun by omega. + exact Hrun. +(* - apply Hstep. + - replace bound with (start + (bound - start)) in Hrun by omega. + exact Hrun. *) +Qed. + +Fixpoint fold_CCcall_spec_ocond {tmp tps} es + : @fold_synth_expr_wellformed tmp tps es -> + OProp3 (HList tp_ft tps) (spec_env_t tmp) GetHighData -> + OProp2 (spec_env_t tmp) GetHighData := + match es with + | HNil => fun _ acc => ounlift32a HNil acc + | HCons x ls e es => + fun wf acc => + let (wf, IHwf) := wf in + fold_CCcall_spec_ocond es IHwf + (omap3 (fun p args se m => + p (HCons x _ (synth_expr_spec me _ e wf se) args) se m) + acc) + end. + +Lemma exists_impl {A} (P:A->Prop) (Q:A->Prop): + (forall x, Q x -> P x) -> + (exists x, Q x) -> + exists x, P x. +Proof. + intros Himpl [x HQ]. + exists x. apply Himpl. exact HQ. +Qed. + +Lemma exists_impl2 {A B} (P:A->B->Prop) (Q:A->B->Prop): + (forall x y, Q x y -> P x y) -> + (exists x y, Q x y) -> + exists x y, P x y. +Proof. + intros Himpl (x & y & HQ). + exists x, y. apply Himpl. exact HQ. +Qed. + +Lemma first_map_osT' (check: Z -> DS bool) start steps m: + (forall i, 0 <= i < steps -> exists v, runStateT (check (start+i)) m = ret (v, m)) -> + exists l, runStateT (first_map check start (start+steps)) m = ret (l, m). +Proof. + intros opt_step. + assert (opt_steps: exists v, + runStateT (oZiteri (fun i l => m <- get;; h <- check (start + i);; put m) steps tt) m = ret (v, m)). + { eapply oZiteri_pure_connect. + intros z z_range l' opt_z. + destruct (opt_step z z_range) as [h opt_check]. + exists tt. + eapply bind_osT. apply get_osT. tauto. + eapply bind_osT. exact opt_check. + apply put_osT. tauto. + } + destruct opt_steps as [v opt_steps]. + exists v. + unfold first_map. + replace (start + steps - start) with steps by omega. + exact opt_steps. +Qed. + +Lemma first_map_osT (check: Z -> DS bool) start bound m: + (forall i, start <= i < bound -> exists v, runStateT (check i) m = ret (v, m)) -> + exists l, runStateT (first_map check start bound) m = ret (l, m). +Proof. + intros Hstep. + replace bound with (start + (bound - start)) by omega. + eapply first_map_osT'. + intros j j_range. + eapply Hstep. omega. +Qed. + +Lemma first_map_osT_weak' (check: Z -> DS bool) start steps m: + (forall i, 0 <= i < steps -> exists v m', runStateT (check (start+i)) m = ret (v, m')) -> + exists l, runStateT (first_map check start (start+steps)) m = ret (l, m). +Proof. + intros opt_step. + assert (opt_steps: exists v, + runStateT (oZiteri (fun i l => m <- get;; h <- check (start + i);; put m) steps tt) m = ret (v, m)). + { eapply oZiteri_pure_connect. + intros z z_range l' opt_z. + destruct (opt_step z z_range) as [h [m' opt_check]]. + exists tt. + eapply bind_osT. apply get_osT. tauto. + eapply bind_osT. exact opt_check. + apply put_osT. tauto. + } + destruct opt_steps as [v opt_steps]. + exists v. + unfold first_map. + replace (start + steps - start) with steps by omega. + exact opt_steps. +Qed. + +Lemma first_map_osT_weak (check: Z -> DS bool) start bound m: + (forall i, start <= i < bound -> exists v m', runStateT (check i) m = ret (v, m')) -> + exists l, runStateT (first_map check start bound) m = ret (l, m). +Proof. + intros Hstep. + replace bound with (start + (bound - start)) by omega. + eapply first_map_osT_weak'. + intros j j_range. + eapply Hstep. omega. +Qed. + + Lemma first_map_mid n (check: Z -> DS bool) start bound: + start <= n < bound -> + forall m v s, runStateT (first_map check start bound) m = ret (v, s) -> + exists v' s', runStateT (check n) m = ret (v', s'). + Proof. + intros n_range. intros m v s opt_map. + unfold first_map in opt_map. + replace (bound - start) with ((n - start) + (bound - n)) in opt_map by omega. + apply oZiteri_mid in opt_map; [|omega|omega]. + destruct opt_map as (v_mid & s_mid & opt_left & opt_right). + assert (m = s_mid). + { apply first_map_pure in opt_left. assumption. } + subst s_mid. + replace (bound - n) with (Z.succ (bound - n - 1)) in opt_right by omega. + apply oZiteri_cut_head in opt_right; [|omega]. + destruct opt_right as (v_l & s_l & opt_n & opt_rest). + rewrite oZiteri_1 in opt_n. + split_run opt_n as (v' & s' & opt_get & opt_n). apply get_osT in opt_get. destruct opt_get; subst v' s'. + split_run opt_n as (v' & s' & opt_n & opt_put). apply put_osT in opt_put. destruct opt_put; subst v_l s_l. + replace (start + (n - start + 0)) with n in opt_n by omega. + exists v', s'. + exact opt_n. + Qed. + + Lemma some_injective_isomorphism {A}{a b : A} : Some a ~~~ Some b -> a ~~~ b. + Proof. + intros some_iso. + split. + - intros F fa. + set (F' x := match x with None => F a | Some x => F x end). + apply (iso_f some_iso (F := F')), fa. + - intros F fb. + set (F' x := match x with None => F a | Some x => F x end). + apply (iso_g some_iso (F := F')), fb. + Qed. + + Lemma set_lenv_cond tmp (se : SpecTree.t tmp) le i htp f v : + let set_se := SpecTree.set i htp f se in + let set_le := PTree.set i v le in + ht_ft_cond f -> + ht_rel f v -> + lenv_cond se le -> + lenv_cond set_se set_le. + Proof. + intros set_se set_le rel f_cond le_cond i' tp' hti'. + set (htp' := @mk_hyper_type_pair tp' hti'). + intros eq; generalize eq. (* save it so that we can change it *) + unfold set_se, set_le. + + destruct (peq i' i) as [ i_eq | i_ne ]. + - rewrite i_eq in eq |- *; rewrite AList.gss in eq. + intros eq'. + rewrite (SpecTree.gess' i _ se _ eq), PTree.gss. + generalize eq. + set (F x := forall eq0 : Some htp ~~~ Some x, + ht_ft_cond (eq0 SpecTree.OptionalType f) /\ + ht_rel_some (eq0 SpecTree.OptionalType f) (Some v)). + apply (some_injective_isomorphism eq F). + intros eq_htp. + rewrite <- functor_identity_paramatricity. + split; [| constructor ]; assumption. + - rewrite AList.gso in eq; try assumption. + intros eq'. + rewrite (SpecTree.geso i' i _ _ eq' eq i_ne), PTree.gso; try assumption. + apply le_cond. + Qed. + + Lemma unset_lenv_cond {tmp i htp f}{se : SpecTree.t tmp}{le} : + (tmp ! i ~~~ None)%alist -> + lenv_cond (SpecTree.set i htp f se) le -> lenv_cond se le. + Proof. + intros iso le_cond i' tp' hti'. + set (htp' := mk_hyper_type_pair tp'). + intros eq. + + destruct (peq i' i) as [ i_eq | i_ne ]. + - exfalso. + rewrite i_eq in eq. + set (is_some (o : option hyper_type_pair) := + match o with None => false | _ => true end). + apply (isomorphism_decidable_neq_absurd is_some (!eq @ iso)); + reflexivity. + - assert (iso' : (AList.set i htp tmp) ! i' %alist ~~~ Some htp') + by (rewrite AList.gso; assumption). + specialize (le_cond i' tp' hti' iso'). + rewrite (SpecTree.geso _ _ _ _ iso' eq i_ne) in le_cond. + apply le_cond. + Qed. + + (* The le is allowed to have more mappings than the se. *) + Lemma set_le_lenv_cond tmp (se : SpecTree.t tmp) le i v : + (tmp ! i)%alist ~~~ None -> lenv_cond se le -> lenv_cond se (PTree.set i v le). + Proof. + intros iso le_cond i' tp' hti'. + set (htp' := @mk_hyper_type_pair tp' hti'). + intros eq; generalize eq. (* save it so that we can change it *) + + destruct (peq i' i) as [ i_eq | i_ne ]. + - exfalso. + rewrite i_eq in eq. + set (is_some (o : option hyper_type_pair) := + match o with None => false | _ => true end). + apply (isomorphism_decidable_neq_absurd is_some (!eq @ iso)); + reflexivity. + - intros eq'. + rewrite PTree.gso; try assumption. + apply le_cond. + Qed. + + (* If the le has more mappings than se, then we can extend the se correspondingly. *) + Lemma set_se_lenv_cond tmp (se : SpecTree.t tmp) le i htp f v : + let set_se := SpecTree.set i htp f se in + ht_ft_cond f -> + ht_rel f v -> + le ! i = Some v -> + lenv_cond se le -> + lenv_cond set_se le. + Proof. + intros set_se f_cond rel le_i_v le_cond i' tp' hti'. + set (htp' := @mk_hyper_type_pair tp' hti'). + intros eq; generalize eq. (* save it so that we can change it *) + unfold set_se. + + destruct (peq i' i) as [ i_eq | i_ne ]. + - rewrite i_eq in eq |- *; rewrite AList.gss in eq. + intros eq'. + rewrite (SpecTree.gess' i _ se _ eq), le_i_v. + generalize eq. + set (F x := forall eq0 : Some htp ~~~ Some x, + ht_ft_cond (eq0 SpecTree.OptionalType f) /\ + ht_rel_some (eq0 SpecTree.OptionalType f) (Some v)). + apply (some_injective_isomorphism eq F). + intros eq_htp. + rewrite <- functor_identity_paramatricity. + split; [| constructor ]; assumption. + - rewrite AList.gso in eq by assumption. + intros eq'. + rewrite (SpecTree.geso i' i _ _ eq' eq i_ne); try assumption. + apply le_cond. + Qed. + + Lemma get_neq_idx_neq {tmp i j}{htp : hyper_type_pair} : + (tmp ! i)%alist ~~~ Some htp -> (tmp ! j)%alist ~~~ None -> i <> j. + Proof. + intros iso_i iso_j i_eq. + rewrite i_eq in iso_i. + set (is_some (o : option hyper_type_pair) := + match o with None => false | _ => true end). + apply (isomorphism_decidable_neq_absurd is_some (!iso_i @ iso_j)); + reflexivity. + Qed. + + Definition Z_to_val : Z -> val := @ht_implement tint_Z32 (int_Z_iso_impl _). + + + Lemma min_ex_elim_range {T} P lo hi Pdec (found_branch : Z -> T) not_found_branch : + ~ P lo -> + match LangDef.min_ex P lo hi Pdec with + | inleft (exist n _) => found_branch n + | inright _ => not_found_branch + end = + match LangDef.min_ex P (Z.succ lo) hi Pdec with + | inleft (exist m _) => found_branch m + | inright _ => not_found_branch + end. + Proof. + intros not_Plo. + destruct + (LangDef.min_ex P lo hi Pdec) as [ (n & n_range & Pn & n_least) | no_n ], + (LangDef.min_ex P (Z.succ lo) hi Pdec) + as [ (m & m_range & Pm & m_least) | no_m ]. + - (* n = m *) + destruct (Ztrichotomy n m) as [ lt | [ eql | gt ]]. + + (* if n < m *) + contradiction (m_least n); try split; try assumption. + (* Z.succ lo <= n (excluding lo = n b/c not_Plo) *) + destruct (proj1 (Z.lt_eq_cases _ _) (proj1 n_range)) as [ lt' | eql ]. + * apply Z.le_succ_l, lt'. + * rewrite eql in not_Plo; contradiction (not_Plo Pn). + + (* if n = n *) + f_equal; exact eql. + + (* if n > m *) + apply Z.gt_lt in gt. + contradiction (n_least m); try split; try assumption. + (* lo <= m (weaker than Z.succ lo <= m) *) + apply Z.lt_le_incl, Z.le_succ_l, m_range. + - (* n = default *) + contradiction (no_m n); split; try apply n_range. + (* Z.succ lo <= n (excluding lo = n b/c not_Plo) *) + destruct (proj1 (Z.lt_eq_cases _ _) (proj1 n_range)) as [ lt' | eql ]. + + apply Z.le_succ_l, lt'. + + rewrite eql in not_Plo; contradiction (not_Plo Pn). + - (* default = m *) + contradiction (no_n m); split; try apply m_range. + (* lo <= m (weaker than Z.succ lo <= m) *) + apply Z.lt_le_incl, Z.le_succ_l, m_range. + - reflexivity. + Qed. + + (** [synth_stmt_cond] is the verification condition for + [synth_stmt_spec_opt] to refine the Clight semantics of [synth_stmt_stmt], + i.e. the one that the user has to manually prove. + This is Figure 4.11 in the thesis. + + Naming convention: _cond is verification conditions. + + [forall se mh, oProp2 (synth_stmt_ocond c dest tmp wf) se mh -> + forall B k b, + synth_stmt_spec_opt c dest tmp wf se mh B k = Some b -> + forall le ml ge, [[se and mh corresponding to le, ml, and ge]] -> + exists r, + (forall B' k', synth_stmt_spec_opt c dest tmp wf se mh B' k' = k' r) /\ + exists le' ml', [[r correspond to le' and ml']] /\ + LangDef.exec_stmt m0 ge env le ml (synth_stmt_stmt c dest tmp) + E0 le' ml' Out_normal] + + as shown in [synth_stmt_correct] below. *) + + + Fixpoint fold_expr_constr_list_cond {tmp tps} es + : @fold_synth_expr_wellformed tmp tps es -> + spec_env_t tmp -> Prop := + match es with + | HNil => fun _ _ => True + | HCons x ls e es => fun wf se => + let (wf, IHwf) := wf in + oProp1 (synth_expr_ocond me tmp e wf /\ +(* + ht_valid_ft_ocond .{ synth_expr_spec tmp e wf } +*) +(* The second conjunct is syntax for this: *) + (omap1 (fun (p : unpair_ft (tp_type_pair x) -> Prop) + (y : spec_env_t tmp) => + p + (synth_expr_spec me tmp e wf y)) + ht_valid_ft_ocond) + )%oprop1 se /\ (* NOTE: implement synth_expr_cond? *) + fold_expr_constr_list_cond es IHwf se + end. + + Fixpoint synth_stmt_cond {returns}(c : cmd_constr returns) dest tmp : + synth_stmt_wellformed c dest tmp -> spec_env_t tmp -> GetHighData -> Prop := + match c as t return synth_stmt_wellformed t dest tmp -> _ with + | CCskip => fun _ _ _ => True + | CClet _ tp _ id c1 c2 => + fun (wf : let tmp' := AList.set id (mk_hyper_type_pair tp) tmp in + ((tmp ! id ~~~ None) * dest_not_exist dest tmp') * + (synth_stmt_wellformed c1 id tmp * synth_stmt_wellformed c2 dest tmp')) + se m => + synth_stmt_cond c1 id _ (cadr wf) se m /\ + (forall v m1, runStateT (synth_stmt_spec_opt c1 id tmp (cadr wf) se) m = ret (v, m1) -> + synth_stmt_cond c2 dest _ (cddr wf) + (SpecTree.set id (mk_hyper_type_pair tp) v se) m1) +(* | CClet _ tp _ id c1 c2 => + fun (wf : let tmp' := AList.set id (mk_hyper_type_pair tp) tmp in + ((tmp ! id ~~~ None) * dest_not_exist dest tmp') * + (synth_stmt_wellformed c1 id tmp * synth_stmt_wellformed c2 dest tmp')) + se m => + forall v m1, runStateT (synth_stmt_spec_opt c1 id tmp (cadr wf) se) m = ret (v, m1) -> + synth_stmt_cond c1 id _ (cadr wf) se m /\ + synth_stmt_cond c2 dest _ (cddr wf) + (SpecTree.set id (mk_hyper_type_pair tp) v se) m1 *) + | CCload tp _ e => fun wf se d => + oProp1 (synth_lexpr_ocond me tmp e wf) se /\ + oProp1 + ((synth_lexpr_spec me tmp e wf se).(ltype_set_ocond) /\ + (synth_lexpr_spec me tmp e wf se).(ltype_get_extra_ocond))%oprop1 d + | CCstore _ _ el er => fun wf se d => + oProp1 (synth_lexpr_ocond me tmp el (car wf)) se /\ + oProp1 (synth_expr_ocond me tmp er (cdr wf)) se /\ + ht_valid_ft_cond (synth_expr_spec me tmp er (cdr wf) se) /\ + oProp1 (synth_lexpr_spec me tmp el (car wf) se).(ltype_set_ocond) d + | CCsequence _ c1 c2 => + fun (wf : synth_stmt_wellformed c1 dest tmp * synth_stmt_wellformed c2 dest tmp) + se m => + synth_stmt_cond c1 dest tmp (car wf) se m /\ + (forall m1, execStateT (synth_stmt_spec_opt c1 dest tmp (car wf) se) m = ret m1 -> + synth_stmt_cond c2 dest tmp (cdr wf) se m1) +(* | CCsequence _ c1 c2 => + fun (wf : synth_stmt_wellformed c1 dest tmp * synth_stmt_wellformed c2 dest tmp) + se m => + forall m1, execStateT (synth_stmt_spec_opt c1 dest tmp (car wf) se) m = ret m1 -> + synth_stmt_cond c1 dest tmp (car wf) se m /\ + synth_stmt_cond c2 dest tmp (cdr wf) se m1 *) + | CCifthenelse _ e c_true c_false => + fun (wf : synth_expr_wellformed tmp e * + (synth_stmt_wellformed c_true dest tmp * synth_stmt_wellformed c_false dest tmp)) + se m => + oProp1 (synth_expr_ocond me tmp e (car wf)) se /\ + (synth_expr_spec me tmp e (car wf) se = true -> + synth_stmt_cond c_true dest tmp (cadr wf) se m) /\ + (synth_expr_spec me tmp e (car wf) se = false -> + synth_stmt_cond c_false dest tmp (cddr wf) se m) + | CCfor id_it id_end e1 e2 c3 => + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + dest_not_exist dest tmp'')) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + synth_stmt_wellformed c3 dest tmp''))) + se m => + oProp1 (synth_expr_ocond me tmp e1 (cadr wf)) se /\ + oProp1 (synth_expr_ocond me tmp e2 (caddr wf)) se /\ + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + forall n, start <= n < bound -> + forall m', + execStateT (Ziteri + (for_step_opt initial_se start (synth_stmt_spec_opt c3 dest _ (cdddr wf))) + (n - start) + (ret tt)) m = ret m' -> + synth_stmt_cond c3 dest _ (cdddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m') + | CCfold tp _ id_it id_end id_recur id_dest e1 e2 e3 c4 => + let htp := mk_hyper_type_pair tp in + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + let tmp''' := AList.set id_recur (mk_hyper_type_pair tp) tmp'' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + ((tmp'' ! id_recur ~~~ None) * ((tmp''' ! id_dest ~~~ None) * + (dest_not_exist dest tmp''' * (synth_stmt_pure c4 ~~~ true)))))) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + (synth_expr_wellformed tmp e3 * synth_stmt_wellformed c4 id_dest tmp''')))) + se m => + oProp1 (synth_expr_ocond me tmp e1 (cadr wf)) se /\ + oProp1 (synth_expr_ocond me tmp e2 (caddr wf)) se /\ + oProp1 (synth_expr_ocond me tmp e3 (cadddr wf)) se /\ + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + let init := synth_expr_spec me tmp e3 (cadddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + forall n, start <= n < bound -> + forall recur, + evalStateT (oZiteri + (fold_step_opt initial_se start (synth_stmt_spec_opt c4 id_dest _ (cddddr wf))) + (n - start) + init) m = ret recur -> + synth_stmt_cond c4 id_dest _ (cddddr wf) + (SpecTree.set id_recur htp recur + (SpecTree.set id_it int_Z32_pair n initial_se)) + m) + | CCcall argt ret prim args => fun wf se _ => + fold_expr_constr_list_cond args wf se + (* | CCcall_ext argt ret prim args => fun wf se _ => + fold_expr_constr_list_cond args wf se *) + | CCyield tp _ e => fun wf se _ => oProp1 (synth_expr_ocond me tmp e wf) se + | CCconstr _ _ _ _ el flds _ => + fun (wf : synth_lexpr_wellformed tmp el * fold_synth_expr_wellformed tmp flds) + se m => + oProp1 (synth_lexpr_ocond me tmp el (car wf)) se /\ + oProp1 (synth_lexpr_spec me tmp el (car wf) se).(ltype_set_ocond) m /\ + fold_expr_constr_list_cond flds (cdr wf) se + | CCassert c => fun wf se m => synth_stmt_cond c dest tmp (cdr wf) se m + | CCdeny c => fun wf se m => synth_stmt_cond c dest tmp (cdr wf) se m + (* | CCghost _ => fun _ _ _ => True *) + | CCpanic _ _ => fun _ _ _ => True + | CCfirst _ id_it id_end id_dest e1 e2 c3 c4 c5 => + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + ((tmp'' ! id_dest ~~~ None) * (dest_not_exist dest tmp'' * + ((synth_stmt_pure c3 ~~~ true)))))) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + (synth_stmt_wellformed c3 id_dest tmp'' * + (synth_stmt_wellformed c4 dest tmp'' * + synth_stmt_wellformed c5 dest tmp''))))) + se m => + oProp1 (synth_expr_ocond me tmp e1 (cadr wf)) se /\ + oProp1 (synth_expr_ocond me tmp e2 (caddr wf)) se /\ + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in +(* let P m n := + evalStateT (synth_stmt_spec_opt c3 id_dest _ (cadddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se))) m + = Some true +(* is_true (synth_stmt_spec_ret id_dest _ _ (synth_stmt_spec_opt c3) (cadddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m) *) + in *) + (forall n, start <= n < bound -> + (* ???: /\ (forall n', start <= n' < n -> ~ P m n') *) + synth_stmt_cond c3 id_dest _ (cadddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m) /\ + (forall n, start <= n < bound (* /\ P m n /\ + (forall n', start <= n' < n -> ~ P m n') *) -> + synth_stmt_cond c4 dest _ (caddddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m) /\ + ((* (forall n, start <= n < bound -> ~ P m n) -> *) + synth_stmt_cond c5 dest _ (cdddddr wf) + (SpecTree.set id_it int_Z32_pair bound + (SpecTree.set id_end int_Z32_pair bound se)) m)) + | CCrespec _ tmp' c spec => + fun (wf : ((tmp ~~~ tmp') * (synth_stmt_pure c ~~~ true)) * + synth_stmt_wellformed c dest tmp) => + synth_stmt_cond c dest tmp (cdr wf) + | CCrespec_opt _ tmp' c spec => + fun (wf : (tmp ~~~ tmp') * synth_stmt_wellformed c dest tmp) => + synth_stmt_cond c dest tmp (cdr wf) + end %alist. + + Inductive synth_stmt_CD (dest: positive) (tmp: AList.t hyper_type_pair) (se: spec_env_t tmp) (m: GetHighData): + forall {returns} (c: cmd_constr returns) (wf: synth_stmt_wellformed c dest tmp), Prop := + | CDskip: forall wf, synth_stmt_CD dest tmp se m CCskip wf + | CDlet: forall r tp hti id c1 c2 wf, + synth_stmt_CD id tmp se m c1 (cadr wf) -> + (forall v m1, runStateT (synth_stmt_spec_opt c1 id tmp (cadr wf) se) m = ret (v, m1) -> + synth_stmt_CD dest _ (SpecTree.set id (mk_hyper_type_pair tp) v se) m1 c2 (cddr wf)) -> + synth_stmt_CD dest tmp se m (@CClet _ r tp hti id c1 c2) wf + | CDload: forall tp hti e wf, + oProp1 (synth_lexpr_ocond me tmp e wf) se -> + oProp1 + (ltype_set_ocond (synth_lexpr_spec me tmp e wf se) /\ + ltype_get_extra_ocond (synth_lexpr_spec me tmp e wf se))%oprop1 m -> + synth_stmt_CD dest tmp se m (@CCload _ tp hti e) wf + | CDstore: forall tp hti el er wf, + oProp1 (synth_lexpr_ocond me tmp el (car wf)) se -> + oProp1 (synth_expr_ocond me tmp er (cdr wf)) se -> + ht_valid_ft_cond (synth_expr_spec me tmp er (cdr wf) se) -> + oProp1 (ltype_set_ocond (synth_lexpr_spec me tmp el (car wf) se)) m -> + synth_stmt_CD dest tmp se m (@CCstore _ tp hti el er) wf + | CDsequence: forall r c1 c2 wf, + synth_stmt_CD dest tmp se m c1 (car wf) -> + (forall m1 v1, runStateT (synth_stmt_spec_opt c1 dest tmp (car wf) se) m = ret (v1, m1) -> + synth_stmt_CD dest tmp se m1 c2 (cdr wf)) -> + synth_stmt_CD dest tmp se m (@CCsequence _ r c1 c2) wf + | CDifthenelse: forall r (e: @expr_constr _ tint_bool _) c_true c_false wf, + oProp1 (synth_expr_ocond me tmp e (car wf)) se -> + (synth_expr_spec me tmp e (car wf) se = true -> synth_stmt_CD dest tmp se m c_true (cadr wf)) -> + (synth_expr_spec me tmp e (car wf) se = false -> synth_stmt_CD dest tmp se m c_false (cddr wf)) -> + synth_stmt_CD dest tmp se m (@CCifthenelse _ r e c_true c_false) wf + | CDfor: forall id_it id_end + (e1 e2: @expr_constr _ tint_Z32 _) (c3: @cmd_constr _ void_unit_pair) wf, + oProp1 (synth_expr_ocond me tmp e1 (cadr wf)) se -> + oProp1 (synth_expr_ocond me tmp e2 (caddr wf)) se -> + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + forall n, start <= n < bound -> + forall v m', + runStateT + (Ziteri + (for_step_opt initial_se start + (synth_stmt_spec_opt c3 dest + (AList.set id_it int_Z32_pair (AList.set id_end int_Z32_pair tmp)) + (cdddr wf))) (n - start) (ret tt)) m = + ret (v, m') -> + synth_stmt_CD dest _ + (SpecTree.set id_it int_Z32_pair n (SpecTree.set id_end int_Z32_pair bound se)) + m' c3 (cdddr wf)) -> + synth_stmt_CD dest tmp se m (@CCfor _ id_it id_end e1 e2 c3) wf + + | CDfirst: forall r id_it id_end id_dest + (e1 e2: @expr_constr _ tint_Z32 _) c3 (c4 c5: @cmd_constr _ r) wf, + oProp1 (synth_expr_ocond me tmp e1 (cadr wf)) se -> + oProp1 (synth_expr_ocond me tmp e2 (caddr wf)) se -> + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + forall n, + start <= n < bound -> + synth_stmt_CD id_dest _ + (SpecTree.set id_it int_Z32_pair n (SpecTree.set id_end int_Z32_pair bound se)) + m c3 (cadddr wf)) -> + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + forall n, + start <= n < bound -> + synth_stmt_CD dest _ + (SpecTree.set id_it int_Z32_pair n (SpecTree.set id_end int_Z32_pair bound se)) + m c4 (caddddr wf)) -> + (let bound := synth_expr_spec me tmp e2 (caddr wf) se in + synth_stmt_CD dest _ + (SpecTree.set id_it int_Z32_pair bound (SpecTree.set id_end int_Z32_pair bound se)) + m c5 (cdddddr wf)) -> + synth_stmt_CD dest tmp se m (@CCfirst _ r id_it id_end id_dest e1 e2 c3 c4 c5) wf + + | CDfold: forall tp hti id_it id_end id_recur id_dest + (e1 e2: @expr_constr _ tint_Z32 _) e3 (c4: @cmd_constr _ (mk_hyper_type_pair tp)) wf, + oProp1 (synth_expr_ocond me tmp e1 (cadr wf)) se -> + oProp1 (synth_expr_ocond me tmp e2 (caddr wf)) se -> + oProp1 (synth_expr_ocond me tmp e3 (cadddr wf)) se -> + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + let init := synth_expr_spec me tmp e3 (cadddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + forall n : Z, + start <= n < bound -> + forall recur m', + runStateT + (oZiteri + (fold_step_opt initial_se start + (synth_stmt_spec_opt c4 id_dest + (AList.set id_recur (mk_hyper_type_pair tp) + (AList.set id_it int_Z32_pair (AList.set id_end int_Z32_pair tmp))) + (cddddr wf))) (n - start) init) m = + ret (recur, m') -> + synth_stmt_CD id_dest _ + (SpecTree.set id_recur _ recur (SpecTree.set id_it int_Z32_pair n initial_se)) + m c4 (cddddr wf)) -> + synth_stmt_CD dest tmp se m (@CCfold _ tp hti id_it id_end id_recur id_dest e1 e2 e3 c4) wf + | CDcall: forall argts r prim args wf, + fold_expr_constr_list_cond args wf se -> + synth_stmt_CD dest tmp se m (@CCcall _ argts r prim args) wf + (* | CDcall_ext: forall argts r addr prim args wf, + fold_expr_constr_list_cond args wf se -> + synth_stmt_CD dest tmp se m (@CCcall_ext _ argts r addr prim args) wf *) + | CDyield: forall tp hti e wf, + oProp1 (synth_expr_ocond me tmp e wf) se -> + synth_stmt_CD dest tmp se m (@CCyield _ tp hti e) wf + | CDconstr: forall tp hti fld_ids fld_tps el flds constr wf, + oProp1 (synth_lexpr_ocond me tmp el (car wf)) se -> + oProp1 (ltype_set_ocond (synth_lexpr_spec me tmp el (car wf) se)) m -> + fold_expr_constr_list_cond flds (cdr wf) se -> + synth_stmt_CD dest tmp se m (@CCconstr _ tp hti fld_ids fld_tps el flds constr) wf + | CDassert: forall c wf, + synth_stmt_CD dest tmp se m c (cdr wf) -> + synth_stmt_CD dest tmp se m (@CCassert _ c) wf + | CDdeny: forall c wf, + synth_stmt_CD dest tmp se m c (cdr wf) -> + synth_stmt_CD dest tmp se m (@CCdeny _ c) wf + (* | CDghost: forall c wf, synth_stmt_CD dest tmp se m (@CCghost _ c) wf *) + | CDpanic: forall tp hti wf, synth_stmt_CD dest tmp se m (@CCpanic _ tp hti) wf + | CDrespec: forall r tmp' (c: cmd_constr r) (spec: spec_env_t tmp' -> DS (tp_ft r)) (wf : (tmp ~~~ tmp') * (synth_stmt_pure c ~~~ true) * synth_stmt_wellformed c dest tmp), + synth_stmt_CD dest tmp se m c (cdr wf) -> + synth_stmt_CD dest tmp se m (@CCrespec _ r tmp' c spec) wf + | CDrespec_opt: forall r tmp' (c: cmd_constr r) (spec: machine_env GetHighData -> spec_env_t tmp' -> DS (tp_ft r)) (wf: (tmp ~~~ tmp') * synth_stmt_wellformed c dest tmp), + synth_stmt_CD dest tmp se m c (cdr wf) -> + synth_stmt_CD dest tmp se m (@CCrespec_opt _ r tmp' c spec) wf + . + + Lemma synth_stmt_CD_cond r (c: cmd_constr r) dest tmp + (wf: synth_stmt_wellformed c dest tmp) (se: spec_env_t tmp) m: + synth_stmt_CD dest tmp se m c wf -> + synth_stmt_cond c dest tmp wf se m. + Proof. + intros CD. + induction CD; simpl; eauto. + - (* CDsequence *) + split. eauto. + intros m1 opt. + exec2run opt as v1. eauto. + - (* CDfor *) + repeat split; eauto. + intros n n_bound m' opt. + exec2run opt as v. eauto. + - (* CDfold *) + repeat split; eauto. + intros n n_bound recur opt. + eval2run opt as m'. eauto. + Qed. + + (* The synth_stmt_obligation is the part of the verification condition to do with CCrespec. + This could be merged into the definition of _cond above, but we treat it seperately + because the respecs are generated by Edsger, so in future works Edsger could maybe also + generate the respec proofs. + *) + + Fixpoint synth_stmt_obligation {returns}(c : cmd_constr returns) dest tmp : + synth_stmt_wellformed c dest tmp -> spec_env_t tmp -> GetHighData -> Prop := + match c as t return synth_stmt_wellformed t dest tmp -> _ with + | CCskip => fun _ _ _ => True + | CClet _ tp _ id c1 c2 => + fun (wf : let tmp' := AList.set id (mk_hyper_type_pair tp) tmp in + ((tmp ! id ~~~ None) * dest_not_exist dest tmp') * + (synth_stmt_wellformed c1 id tmp * synth_stmt_wellformed c2 dest tmp')) + se m => + synth_stmt_obligation c1 id _ (cadr wf) se m /\ + (forall v m', runStateT (synth_stmt_spec_opt c1 id tmp (cadr wf) se) m = ret (v, m') -> + synth_stmt_obligation c2 dest _ (cddr wf) (SpecTree.set id (mk_hyper_type_pair tp) v se) m') + | CCload tp _ e => fun wf se _ => True + | CCstore _ _ el er => fun _ _ _ => True + | CCsequence _ c1 c2 => + fun (wf : synth_stmt_wellformed c1 dest tmp * synth_stmt_wellformed c2 dest tmp) + se m => + synth_stmt_obligation c1 dest tmp (car wf) se m /\ + (forall m', execStateT (synth_stmt_spec_opt c1 dest tmp (car wf) se) m = ret m' -> + synth_stmt_obligation c2 dest tmp (cdr wf) se m') + | CCifthenelse _ e c_true c_false => + fun (wf : synth_expr_wellformed tmp e * + (synth_stmt_wellformed c_true dest tmp * synth_stmt_wellformed c_false dest tmp)) + se m => + (synth_expr_spec me tmp e (car wf) se = true -> + synth_stmt_obligation c_true dest tmp (cadr wf) se m) /\ + (synth_expr_spec me tmp e (car wf) se = false -> + synth_stmt_obligation c_false dest tmp (cddr wf) se m) + | CCfor id_it id_end e1 e2 c3 => + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + dest_not_exist dest tmp'')) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + synth_stmt_wellformed c3 dest tmp''))) + se m => + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + forall n, start <= n < bound -> + forall m', execStateT (Ziteri + (for_step_opt initial_se start (synth_stmt_spec_opt c3 dest _ (cdddr wf))) + (n - start) + (ret tt)) m = ret m' -> + synth_stmt_obligation c3 dest _ (cdddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m') + | CCfold tp _ id_it id_end id_recur id_dest e1 e2 e3 c4 => + let htp := mk_hyper_type_pair tp in + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + let tmp''' := AList.set id_recur (mk_hyper_type_pair tp) tmp'' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + ((tmp'' ! id_recur ~~~ None) * ((tmp''' ! id_dest ~~~ None) * + (dest_not_exist dest tmp''' * (synth_stmt_pure c4 ~~~ true)))))) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + (synth_expr_wellformed tmp e3 * synth_stmt_wellformed c4 id_dest tmp''')))) + se m => + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + let init := synth_expr_spec me tmp e3 (cadddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + forall n, start <= n < bound -> + forall recur, + evalStateT (oZiteri + (fold_step_opt initial_se start (synth_stmt_spec_opt c4 id_dest _ (cddddr wf))) + (n - start) + init) m = ret recur -> + synth_stmt_obligation c4 id_dest _ (cddddr wf) + (SpecTree.set id_recur htp recur + (SpecTree.set id_it int_Z32_pair n initial_se)) + m) + | CCcall argt ret prim args => fun _ _ _ => True + (* | CCcall_ext argt ret addr prim args => fun _ _ _ _ => True *) + | CCyield tp _ e => fun _ _ _ => True + | CCconstr _ _ _ _ el flds _ => + fun (wf : synth_lexpr_wellformed tmp el * fold_synth_expr_wellformed tmp flds) + se m => True + | CCassert c => + fun (wf: (synth_stmt_pure c ~~~ true) * + synth_stmt_wellformed c dest tmp) se m => + synth_stmt_obligation c dest tmp (cdr wf) se m + | CCdeny c => + fun (wf: (synth_stmt_pure c ~~~ true) * + synth_stmt_wellformed c dest tmp) se m => + synth_stmt_obligation c dest tmp (cdr wf) se m + (* | CCghost c => + fun (wf: synth_stmt_wellformed c dest tmp) se m => + synth_stmt_obligation c dest tmp wf se m *) + | CCpanic _ _ => fun _ _ _ => True + | CCfirst _ id_it id_end id_dest e1 e2 c3 c4 c5 => + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + ((tmp'' ! id_dest ~~~ None) * (dest_not_exist dest tmp'' * + ((synth_stmt_pure c3 ~~~ true)))))) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + (synth_stmt_wellformed c3 id_dest tmp'' * + (synth_stmt_wellformed c4 dest tmp'' * + synth_stmt_wellformed c5 dest tmp''))))) + se m => + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in +(* let P m n := + evalStateT (synth_stmt_spec_opt c3 id_dest _ (cadddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se))) m + = Some true +(* is_true (synth_stmt_spec_ret id_dest _ _ (synth_stmt_spec_opt c3) (cadddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m) *) + in *) + (forall n, start <= n < bound -> + (* ???: /\ (forall n', start <= n' < n -> ~ P m n') *) + synth_stmt_obligation c3 id_dest _ (cadddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m) /\ + (forall n, start <= n < bound (* /\ P m n /\ + (forall n', start <= n' < n -> ~ P m n') *) -> + synth_stmt_obligation c4 dest _ (caddddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m) /\ + ((* (forall n, start <= n < bound -> ~ P m n) -> *) + synth_stmt_obligation c5 dest _ (cdddddr wf) + (SpecTree.set id_it int_Z32_pair bound + (SpecTree.set id_end int_Z32_pair bound se)) m)) + | CCrespec _ tmp' c spec => + fun (wf : ((tmp ~~~ tmp') * (synth_stmt_pure c ~~~ true)) * + synth_stmt_wellformed c dest tmp) se m => + (forall se m, runStateT (synth_stmt_spec_opt c dest tmp (cdr wf) se) m = + runStateT (spec (iso_f (caar wf) se)) m) /\ + synth_stmt_obligation c dest _ (cdr wf) se m + | CCrespec_opt _ tmp' c spec => + fun (wf : (tmp ~~~ tmp') * synth_stmt_wellformed c dest tmp) se m => + (forall se m, runStateT (spec me (iso_f (car wf) se)) m = + runStateT (synth_stmt_spec_opt c dest _ (cdr wf) se) m) /\ + (* synth_stmt_side_cond c dest _ (cdr wf) se m /\ *) + synth_stmt_obligation c dest _ (cdr wf) se m + end %alist. + + Inductive synth_stmt_RC (dest:positive) (tmp: AList.t hyper_type_pair) (se: spec_env_t tmp) (m: GetHighData): + forall {returns} (c: cmd_constr returns) (wf: synth_stmt_wellformed c dest tmp), Prop := + | RCskip: forall wf, + synth_stmt_RC dest tmp se m CCskip wf + | RClet: forall r tp hti id c1 c2 wf, + synth_stmt_RC id _ se m c1 (cadr wf) -> + (forall v m', + runStateT (synth_stmt_spec_opt c1 id tmp (cadr wf) se) m = ret (v, m') -> + synth_stmt_RC dest _ (SpecTree.set id (mk_hyper_type_pair tp) v se) m' c2 (cddr wf)) -> + synth_stmt_RC dest tmp se m (@CClet _ r tp hti id c1 c2) wf + | RCload: forall tp hti e wf, + ht_ft_cond (ltype_get (synth_lexpr_spec me tmp e wf se) m) -> + synth_stmt_RC dest tmp se m (@CCload _ tp hti e) wf + | RCstore: forall tp hti el er wf, + synth_stmt_RC dest tmp se m (@CCstore _ tp hti el er) wf + | RCsequence: forall r c1 c2 wf, + (forall v' m', runStateT (synth_stmt_spec_opt c1 dest tmp (car wf) se) m = ret (v', m') -> + synth_stmt_RC dest tmp se m' c2 (cdr wf)) -> + synth_stmt_RC dest tmp se m (@CCsequence _ r c1 c2) wf + | RCifthenelse: forall r (e: @expr_constr _ tint_bool _) c_true c_false wf, + (synth_expr_spec me tmp e (car wf) se = true -> + synth_stmt_RC dest tmp se m c_true (cadr wf)) -> + (synth_expr_spec me tmp e (car wf) se = false -> + synth_stmt_RC dest tmp se m c_false (cddr wf)) -> + synth_stmt_RC dest tmp se m (@CCifthenelse _ r e c_true c_false) wf + | RCfor: forall id_it id_end + (e1 e2: @expr_constr _ tint_Z32 _) (c3: @cmd_constr _ void_unit_pair) wf, +(* (let start := synth_expr_spec tmp e1 (cadr wf) se in + let bound := synth_expr_spec tmp e2 (caddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + forall n, start <= n < bound -> + forall m', execStateT (Ziteri + (for_step_opt initial_se start (synth_stmt_spec_opt c3 dest _ (cdddr wf))) + (n - start) + (ret tt)) m = ret m' -> + synth_stmt_RC dest _ + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) + m' c3 (cdddr wf)) -> *) + synth_stmt_RC dest tmp se m (@CCfor _ id_it id_end e1 e2 c3) wf + + | RCfirst: forall r id_it id_end id_dest + (e1 e2: @expr_constr _ tint_Z32 _) c3 (c4 c5: @cmd_constr _ r) wf, + oProp1 (synth_expr_ocond me tmp e1 (cadr wf)) se -> + oProp1 (synth_expr_ocond me tmp e2 (caddr wf)) se -> + let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + (forall n, start <= n < bound (* /\ P m n /\ + (forall n', start <= n' < n -> ~ P m n') *) -> + synth_stmt_RC dest _ + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) + m c4 (caddddr wf)) -> + ((* (forall n, start <= n < bound -> ~ P m n) -> *) + synth_stmt_RC dest _ + (SpecTree.set id_it int_Z32_pair bound + (SpecTree.set id_end int_Z32_pair bound se)) m c5 (cdddddr wf)) -> + synth_stmt_RC dest tmp se m (@CCfirst _ r id_it id_end id_dest e1 e2 c3 c4 c5) wf + + | RCfold: forall tp hti id_it id_end id_recur id_dest + (e1 e2: @expr_constr _ tint_Z32 _) e3 (c4: @cmd_constr _ (mk_hyper_type_pair tp)) wf, + oProp1 (synth_expr_ocond me tmp e1 (cadr wf)) se -> + oProp1 (synth_expr_ocond me tmp e2 (caddr wf)) se -> + oProp1 (synth_expr_ocond me tmp e3 (cadddr wf)) se -> + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + let init := synth_expr_spec me tmp e3 (cadddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + (* ht_ft_cond *) + forall n, start <= n < bound -> + forall recur, + evalStateT (oZiteri + (fold_step_opt initial_se start (synth_stmt_spec_opt c4 id_dest _ (cddddr wf))) + (n - start) + init) m = ret recur -> + synth_stmt_RC id_dest _ + (SpecTree.set id_recur (mk_hyper_type_pair tp) recur + (SpecTree.set id_it int_Z32_pair n initial_se)) m c4 (cddddr wf)) -> + synth_stmt_RC dest tmp se m (@CCfold _ tp hti id_it id_end id_recur id_dest e1 e2 e3 c4) wf + | RCcall: forall argts r prim args wf, + ht_list_ft_cond (map2_synth_expr_spec args se wf) -> + ht_list_valid_ft_cond (map2_synth_expr_spec args se wf) -> + synth_stmt_RC dest tmp se m (@CCcall _ argts r prim args) wf + (* | RCcall_ext: forall argts r addr prim args wf, + ht_list_ft_cond (map2_synth_expr_spec args se wf) -> + ht_list_valid_ft_cond (map2_synth_expr_spec args se wf) -> + synth_stmt_RC dest tmp se m (@CCcall_ext _ argts r addr prim args) wf *) + | RCyield: forall tp hti e wf, + oProp1 (synth_expr_ocond me tmp e wf) se -> + synth_stmt_RC dest tmp se m (@CCyield _ tp hti e) wf + | RCconstr: forall tp hti fld_ids fld_tps el flds constr wf, + synth_stmt_RC dest tmp se m (@CCconstr _ tp hti fld_ids fld_tps el flds constr) wf + | RCassert: forall c wf, + synth_stmt_RC dest tmp se m (@CCassert _ c) wf + | RCdeny: forall c wf, + synth_stmt_RC dest tmp se m (@CCdeny _ c) wf + | RCghost: forall c wf, + synth_stmt_RC dest tmp se m (@CCdeny _ c) wf + | RCpanic: forall tp hti wf, + synth_stmt_RC dest tmp se m (@CCpanic _ tp hti) wf + | RCrespec: forall r tmp' (c: cmd_constr r) (spec: spec_env_t tmp' -> DS (tp_ft r)) (wf : (tmp ~~~ tmp') * (synth_stmt_pure c ~~~ true) * synth_stmt_wellformed c dest tmp), + (forall v m', runStateT (spec (iso_f (caar wf) se)) m = ret (v, m') -> ht_ft_cond v) -> +(* + (runStateT (synth_stmt_spec_opt c dest tmp (cdr wf) se) m = + runStateT (spec (iso_f (caar wf) se)) m) -> + synth_stmt_RC dest _ se m c (cdr wf) -> +*) + synth_stmt_RC dest tmp se m (@CCrespec _ r tmp' c spec) wf + | RCrespec_opt: forall r tmp' (c: cmd_constr r) (spec: machine_env GetHighData -> spec_env_t tmp' -> DS (tp_ft r)) (wf: (tmp ~~~ tmp') * synth_stmt_wellformed c dest tmp), + (forall v m', runStateT (spec me (iso_f (car wf) se)) m = ret (v, m') -> ht_ft_cond v) -> +(* + (runStateT (synth_stmt_spec_opt c dest _ (cdr wf) se) m = + runStateT (spec (iso_f (car wf) se)) m) -> + synth_stmt_RC dest _ se m c (cdr wf) -> +*) + synth_stmt_RC dest tmp se m (@CCrespec_opt _ r tmp' c spec) wf + . + + Fixpoint synth_stmt_ret_cond {returns}(c : cmd_constr returns) dest tmp : + synth_stmt_wellformed c dest tmp -> spec_env_t tmp -> GetHighData -> Prop := + match c as t return synth_stmt_wellformed t dest tmp -> _ with + | CCskip => fun _ _ _ => True + | CClet _ tp _ id c1 c2 => + fun (wf : let tmp' := AList.set id (mk_hyper_type_pair tp) tmp in + ((tmp ! id ~~~ None) * dest_not_exist dest tmp') * + (synth_stmt_wellformed c1 id tmp * synth_stmt_wellformed c2 dest tmp')) + se m => + synth_stmt_ret_cond c1 id _ (cadr wf) se m /\ + forall v m', runStateT (synth_stmt_spec_opt c1 id tmp (cadr wf) se) m = ret (v, m') -> + synth_stmt_ret_cond c2 dest _ (cddr wf) (SpecTree.set id (mk_hyper_type_pair tp) v se) m' + (* let p1 := synth_stmt_ret_cond c1 id _ (cadr wf) se m in + let p2 := synth_stmt_ret_cond c2 dest _ (cddr wf) (SpecTree.set id (mk_hyper_type_pair tp) v se) m' in + (p1 /\ p2) *) + | CCload tp _ e => fun wf se m => + ht_ft_cond (ltype_get (synth_lexpr_spec me tmp e wf se) m) + | CCstore _ _ el er => fun _ _ _ => True + | CCsequence _ c1 c2 => + fun (wf : synth_stmt_wellformed c1 dest tmp * synth_stmt_wellformed c2 dest tmp) + se m => + forall m', execStateT (synth_stmt_spec_opt c1 dest tmp (car wf) se) m = ret m' -> + synth_stmt_ret_cond c2 dest tmp (cdr wf) se m' + (* + let p1 := synth_stmt_ret_cond c1 dest tmp (car wf) se m in + let p2 := synth_stmt_ret_cond c2 dest tmp (cdr wf) se m' in + (p1 /\ p2) *) + | CCifthenelse _ e c_true c_false => + fun (wf : synth_expr_wellformed tmp e * + (synth_stmt_wellformed c_true dest tmp * synth_stmt_wellformed c_false dest tmp)) + se m => + let obt := synth_stmt_ret_cond c_true dest tmp (cadr wf) se m in + let obf := synth_stmt_ret_cond c_false dest tmp (cddr wf) se m in + (synth_expr_spec me tmp e (car wf) se = true -> obt) /\ + (synth_expr_spec me tmp e (car wf) se = false -> obf) + | CCfor id_it id_end e1 e2 c3 => + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + dest_not_exist dest tmp'')) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + synth_stmt_wellformed c3 dest tmp''))) + se m => True + (* + (let start := synth_expr_spec tmp e1 (cadr wf) se in + let bound := synth_expr_spec tmp e2 (caddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + forall n, start <= n < bound -> + forall m', execStateT (Ziteri + (for_step_opt initial_se start (synth_stmt_spec_opt c3 dest _ (cdddr wf))) + (n - start) + (ret tt)) m = ret m' -> + synth_stmt_ret_cond c3 dest _ (cdddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m') *) + | CCfold tp _ id_it id_end id_recur id_dest e1 e2 e3 c4 => + let htp := mk_hyper_type_pair tp in + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + let tmp''' := AList.set id_recur (mk_hyper_type_pair tp) tmp'' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + ((tmp'' ! id_recur ~~~ None) * ((tmp''' ! id_dest ~~~ None) * + (dest_not_exist dest tmp''' * (synth_stmt_pure c4 ~~~ true)))))) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + (synth_expr_wellformed tmp e3 * synth_stmt_wellformed c4 id_dest tmp''')))) + se m => + oProp1 (synth_expr_ocond me tmp e1 (cadr wf)) se /\ + oProp1 (synth_expr_ocond me tmp e2 (caddr wf)) se /\ + oProp1 (synth_expr_ocond me tmp e3 (cadddr wf)) se /\ + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + let init := synth_expr_spec me tmp e3 (cadddr wf) se in + let initial_se := SpecTree.set id_end int_Z32_pair bound se in + (* ht_ft_cond *) + forall n, start <= n < bound -> + forall recur, + evalStateT (oZiteri + (fold_step_opt initial_se start (synth_stmt_spec_opt c4 id_dest _ (cddddr wf))) + (n - start) + init) m = ret recur -> + synth_stmt_ret_cond c4 id_dest _ (cddddr wf) + (SpecTree.set id_recur htp recur + (SpecTree.set id_it int_Z32_pair n initial_se)) + m) + | CCcall argt ret prim args => fun wf se _ => + (* fold_expr_constr_list_cond args wf se (* /\ senv_cond se *) + /\ *) ht_list_ft_cond (map2_synth_expr_spec args se wf) + /\ ht_list_valid_ft_cond (map2_synth_expr_spec args se wf) + | CCyield tp _ e => fun wf se _ => oProp1 (synth_expr_ocond me tmp e wf) se + | CCconstr _ _ _ _ el flds _ => fun _ _ _ => True + | CCassert c => fun _ _ _ => True + | CCdeny c => fun _ _ _ => True + (* | CCghost c => fun _ _ _ => True *) + | CCpanic _ _ => fun _ _ _ => True + | CCfirst _ id_it id_end id_dest e1 e2 c3 c4 c5 => + fun (wf : let tmp' := AList.set id_end int_Z32_pair tmp in + let tmp'' := AList.set id_it int_Z32_pair tmp' in + ((tmp ! id_end ~~~ None) * ((tmp' ! id_it ~~~ None) * + ((tmp'' ! id_dest ~~~ None) * (dest_not_exist dest tmp'' * + ((synth_stmt_pure c3 ~~~ true)))))) * + (synth_expr_wellformed tmp e1 * (synth_expr_wellformed tmp e2 * + (synth_stmt_wellformed c3 id_dest tmp'' * + (synth_stmt_wellformed c4 dest tmp'' * + synth_stmt_wellformed c5 dest tmp''))))) + se m => + oProp1 (synth_expr_ocond me tmp e1 (cadr wf)) se /\ + oProp1 (synth_expr_ocond me tmp e2 (caddr wf)) se /\ + (let start := synth_expr_spec me tmp e1 (cadr wf) se in + let bound := synth_expr_spec me tmp e2 (caddr wf) se in + (* forall n, start <= n < bound -> + (* ???: /\ (forall n', start <= n' < n -> ~ P m n') *) + synth_stmt_ret_cond c3 id_dest _ (cadddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m *) + (forall n, start <= n < bound (* /\ P m n /\ + (forall n', start <= n' < n -> ~ P m n') *) -> + synth_stmt_ret_cond c4 dest _ (caddddr wf) + (SpecTree.set id_it int_Z32_pair n + (SpecTree.set id_end int_Z32_pair bound se)) m) /\ + ((* (forall n, start <= n < bound -> ~ P m n) -> *) + synth_stmt_ret_cond c5 dest _ (cdddddr wf) + (SpecTree.set id_it int_Z32_pair bound + (SpecTree.set id_end int_Z32_pair bound se)) m)) + | CCrespec _ tmp' c spec => + fun (wf : ((tmp ~~~ tmp') * (synth_stmt_pure c ~~~ true)) * + synth_stmt_wellformed c dest tmp) se m => + (forall v m', runStateT (spec (iso_f (caar wf) se)) m = ret (v, m') -> ht_ft_cond v) +(* + runStateT (synth_stmt_spec_opt c dest tmp (cdr wf) se) m = + runStateT (spec (iso_f (caar wf) se)) m /\ + synth_stmt_ret_cond c dest _ (cdr wf) se m +*) + | CCrespec_opt _ tmp' c spec => + fun (wf : (tmp ~~~ tmp') * synth_stmt_wellformed c dest tmp) se m => + (forall v m', runStateT (spec me (iso_f (car wf) se)) m = ret (v, m') -> ht_ft_cond v) +(* + runStateT (synth_stmt_spec_opt c dest _ (cdr wf) se) m = + runStateT (spec (iso_f (car wf) se)) m /\ + (* synth_stmt_side_cond c dest _ (cdr wf) se m /\ *) + synth_stmt_ret_cond c dest _ (cdr wf) se m +*) + end %alist. + + Lemma synth_stmt_RC_ret_cond r (c: cmd_constr r) dest tmp + (wf: synth_stmt_wellformed c dest tmp) (se: spec_env_t tmp) m: + synth_stmt_RC dest tmp se m c wf -> + synth_stmt_ret_cond c dest tmp wf se m. + Proof. + intros RC. + induction RC; simpl; auto. + intros m' opt_c1. + exec2run opt_c1 as v. + eauto. + Qed. + + Lemma map2_synth_expr_spec_satisfies_ft_cond {tmp argt} + {arg : expr_constr_list argt} + (argc : expr_constr_prf_conj arg) + (wf : fold_synth_expr_wellformed tmp arg) + (se : spec_env_t tmp) : + fold_expr_constr_list_cond arg wf se -> + senv_cond se -> + ht_list_ft_cond (map2_synth_expr_spec arg se wf). + Proof. + intros cd se_cond. + induction arg. + - constructor. + - simpl in argc, wf, cd |- *. + destruct wf as [ wf arg_wf ]. + destruct argc as [ argc [ ht ec ] ]. + destruct cd as [ ecd cd ]. + rewrite oand1_distr in ecd. + destruct ecd as [ ecd e_valid ]. + constructor; + [ apply (synth_expr_spec_satisfies_ft_cond me) + | apply IHarg ]; + assumption. + Qed. + + Lemma map2_synth_expr_spec_satisfies_valid_ft_cond {tmp argt} + {arg : expr_constr_list argt} + (argc : expr_constr_prf_conj arg) + (wf : fold_synth_expr_wellformed tmp arg) + (se : spec_env_t tmp) : + fold_expr_constr_list_cond arg wf se -> + senv_cond se -> + ht_list_valid_ft_cond (map2_synth_expr_spec arg se wf). + Proof. + intros cd se_cond. + induction arg. + - constructor. + - simpl in argc, wf, cd |- *. + destruct wf as [ wf arg_wf ]. + destruct argc as [ argc [ ht ec ] ]. + destruct cd as [ ecd cd ]. + rewrite oand1_distr in ecd. + destruct ecd as [ ecd e_valid ]. + constructor; + [ apply ht_valid_ft_ocond_same; + rewrite OProp1map1 in e_valid by tauto + | apply IHarg ]; + assumption. + Qed. + +Lemma synth_stmt_spec_opt_pure {is_realizing returns htr}(c : cmd_constr returns) : + synth_stmt_pure c ~~~ true -> + cmd_constr_prf is_realizing returns htr c -> + forall dest tmp wf se d v' d', + runStateT (synth_stmt_spec_opt c dest tmp wf se) d = ret (v', d') -> + d = d'. +Proof. + intros X cc. + assert (is_pure := iso_g X (F := fun b => b = true) eq_refl); clear X. + intros dest tmp wf se d v' d' Hrun. + revert dest tmp wf se d v' d' Hrun. + induction cc; try discriminate is_pure; intros; simpl in Hrun. + - (* CCskip *) + apply ret_osT in Hrun. apply Hrun. + - (* CClet *) + split_run Hrun as (v_ & d_ & Hrun1 & Hrun2). + simpl in is_pure. apply andb_prop in is_pure. + transitivity d_. + + eapply IHcc1. apply is_pure. exact Hrun1. + + eapply IHcc2. apply is_pure. exact Hrun2. + - (* CCload *) + apply gets_osT in Hrun. apply Hrun. + - (* CCsequence *) + split_run Hrun as (v_ & d_ & Hrun1 & Hrun2). + simpl in is_pure. apply andb_prop in is_pure. + transitivity d_. + + eapply IHcc1. apply is_pure. exact Hrun1. + + eapply IHcc2. apply is_pure. exact Hrun2. + - (* CCifthenelse *) + simpl in is_pure. apply andb_prop in is_pure. + match type of Hrun with + context [if ?cond then _ else _] => destruct cond + end. + + eapply IHcc1. apply is_pure. exact Hrun. + + eapply IHcc2. apply is_pure. exact Hrun. + - (* CCfirst *) + simpl in is_pure. apply andb_prop in is_pure. + destruct is_pure as [is_pure1 is_pure3]. + apply andb_prop in is_pure1. + destruct is_pure1 as [is_pure1 is_pure2]. + split_run Hrun as (v1 & d1 & Hrun1 & Hrun2). + split_run Hrun2 as (v2 & d2 & Hrun2 & Hrun3). + transitivity d1. + { apply get_osT in Hrun1. apply Hrun1. } + transitivity d2. + { apply first_map_pure in Hrun2. exact Hrun2. } + clear Hrun2. + dest_first Hrun3. + + eapply IHcc2. apply is_pure2. exact Hrun3. + + eapply IHcc3. apply is_pure3. exact Hrun3. + - (* CCfold *) + eapply fold_pure. + 2: exact Hrun. + intros m v_ m_ se_. + apply IHcc. simpl in wf. apply wf. reflexivity. + - (* CCcall *) (* TODO *) + eapply PRIMis_pure. apply is_pure. + apply (runStateT_execStateT Hrun). + (* - CCcall_ext (* TODO *) + eapply PRIMis_pure. apply is_pure. + apply (runStateT_execStateT Hrun). *) + - (* CCyield *) + apply ret_osT in Hrun. apply Hrun. + - (* CCconstr *) + split_run Hrun as (v_ & d_ & Hrun1 & Hrun2). + apply guard_osT in Hrun2. + transitivity d_. + + eapply IHcc. apply is_pure. exact Hrun1. + + apply Hrun2. + - (* CCassert *) + split_run Hrun as (v_ & d_ & Hrun1 & Hrun2). + apply guard_osT in Hrun2. + transitivity d_. + + eapply IHcc. apply is_pure. exact Hrun1. + + apply Hrun2. + (* - (* CCghost *) + eapply IHcc. apply is_pure. exact Hrun. *) + - (* CCPdummy *) + admit. + - (* CCpanic *) + admit. +Admitted. + Lemma set_senv_cond tmp (se : SpecTree.t tmp) i htp v: + ht_ft_cond v -> + senv_cond se -> + senv_cond (SpecTree.set i htp v se). + Proof. + unfold senv_cond. + intros ft_v se_cond. + intros i' tp hti eq. + generalize eq. + destruct (peq i' i) as [i_eq | i_ne]. + - (* i' = i *) + subst i'. rewrite AList.gss in eq. + intros eq'. + rewrite (SpecTree.gess' i _ se _ eq). + generalize eq. + set (F x := forall eq0 : Some htp ~~~ Some x, ht_ft_cond (eq0 SpecTree.OptionalType v)). + apply (some_injective_isomorphism eq F). unfold F. + intros eq_htp. + rewrite <- functor_identity_paramatricity. exact ft_v. + - (* i' != i *) + rewrite AList.gso in eq; [|assumption]. + intros eq'. + rewrite (SpecTree.geso i' i _ _ eq' eq i_ne). + apply se_cond. + Qed. + +Lemma for_mid_opt' {tmp id_it} (se : spec_env_t tmp): + let tmp' := AList.set id_it int_Z32_pair tmp in + forall f s0 start steps s, + steps >= 0 -> + runStateT (Ziteri (@for_step_opt _ id_it se start f) + steps + (ret tt)) s0 = ret (tt, s) -> + forall n, 0%Z <= n <= steps -> + exists s', + runStateT (Ziteri (@for_step_opt _ id_it se start f) n (ret tt)) s0 = ret (tt, s') /\ + runStateT (Ziteri (@for_step_opt _ id_it se (start+n) f) (steps-n) (ret tt)) s' = ret (tt, s). +Proof. + intros tmp' f s0 start steps s steps_range Hrun. + refine (Z_nonneg_peano_ind _ _ _ _). + - (* n = 0 *) + intros n_range. + rewrite Ziteri_base by omega. + exists s0. + split. reflexivity. + replace (start + 0) with start by omega. + replace (steps - 0) with steps by omega. + exact Hrun. + - (* n > 0 *) + intros z z_nonneg Hind Succ_z_range. + assert (z_range: 0 <= z <= steps) by omega. + specialize (Hind z_range). + destruct Hind as (s' & opt_z & opt_rest). + set (rest := (steps - (Z.succ z))). + replace (steps - z) with (Z.succ rest) in opt_rest by (unfold rest; omega). + apply for_cut_opt' in opt_rest; [|unfold rest; omega]. + destruct opt_rest as (s'0 & opt_succ_z & opt_rest). + exists s'0. + split. + + rewrite Ziteri_succ by omega. + unfold for_step_opt at 1. + eapply bind_osT. exact opt_z. + exact opt_succ_z. + + replace (start + Z.succ z) with (start + z + 1) by omega. + exact opt_rest. + - (* n < 0 *) + intros p p_range. + pose (Pos2Z.neg_is_neg p) as n_neg. + omega. +Qed. + + Opaque AList.get. + (* This is the statement of correctness. The following lemmas will prove it for each synthesis function. + + Note that the statement is a little tricky, because it is parametrized by a type of wellformedness condition. + In the cases below, wellformed, spec_cond, ... will be instantiated to synth_stmt_wellformed, synth_stmt_spec_cond, etc. + (See Theorem synth_stmt_spec_correct for an example.) + *) + (* Variable m0 : MemLow. *) +Definition temp_env : Type := PTree.t val. + + Lemma opt_bool_dec_t__f_none (x:option bool): + {x = Some true} + {x = Some false \/ x = None}. + Proof. + destruct x as [b|]; [destruct b|]; tauto. + Qed. + + Lemma Z2Nat_nonpos_zero: + forall n, n <= 0%Z -> Z.to_nat n = 0%nat. + Proof. + destruct n; [reflexivity| |reflexivity]. + specialize (Pos2Z.is_pos p); omega. + Qed. + + Lemma min_ex_empty P lo hi Pn: + lo >= hi -> + exists no_n, LangDef.min_ex P lo hi Pn = inright no_n. + Proof. + intros range. + destruct (LangDef.min_ex P lo hi Pn) as [(n & n_range) | no_n]. + - omega. + - exists no_n. reflexivity. + Qed. + + Lemma min_ex_lo_satisfy (P:Z->Prop) lo hi Pdec + (lo_bound: lo <= lo < hi) + (lo_pass: P lo): + exists Q, LangDef.min_ex P lo hi Pdec = inleft (exist _ lo Q). + Proof. + destruct (LangDef.min_ex P lo hi Pdec) as [[n (n_range & Pn & n_least)] | no_n]. + - assert (n = lo). + { destruct (Z_dec n lo) as [[Hlt | Hgt] | Heq]. + - omega. + - specialize (n_least lo). assert (~ P lo) by (apply n_least; omega). + contradiction. + - assumption. } + subst n. + eexists. reflexivity. + - exfalso. specialize (no_n _ lo_bound). contradiction. + Qed. + + Lemma min_ex_skip_lo {A} {Ql: Z->A} {Qr: A} {P: Z->Prop} {lo hi Pdec}: + ~ P lo -> + match LangDef.min_ex P lo hi Pdec with + | inleft (exist first_n _) => Ql first_n + | inright _ => Qr + end = + match LangDef.min_ex P (Z.succ lo) hi Pdec with + | inleft (exist first_n _) => Ql first_n + | inright _ => Qr + end. + Proof. + intros fail_lo. + destruct (LangDef.min_ex P lo hi Pdec) as [(n & n_range & Pn & n_least)|no_n]; destruct (LangDef.min_ex P (Z.succ lo) hi Pdec) as [(m & m_range & Pm & m_least) | no_m]; [|exfalso|exfalso|]. + - (* Ql n = Ql m *) + destruct (Z_dec n m) as [[Hlt | Hgt] | Heq]. + + (* n < m *) + assert (Z.succ lo <= n). + { destruct (Z_dec lo n) as [[H'lt | H'gt] | H'eq]; [omega|omega|]. + subst n. contradiction. } + specialize (m_least n (conj H Hlt)). contradiction. + + (* n > m *) + exfalso. refine (n_least m _ Pm). + omega. + + (* n = m *) + f_equal. assumption. + - (* Ql n = Qr *) + destruct (Z_dec lo n) as [[Hlt | Hgt] | Heq]. + + (* lo < n *) + refine (no_m n _ Pn). omega. + + (* lo > n *) + omega. + + (* lo = n *) + subst n. contradiction. + - (* Qr = Ql m *) + refine (no_n m _ Pm). omega. + - (* Qr = Qr *) + reflexivity. + Qed. + + Definition synth_stmt_veri_cond + {returns}(c : cmd_constr returns) dest tmp wf se mh := + (* synth_stmt_side_cond c dest tmp wf se mh /\ *) + synth_stmt_cond c dest tmp wf se mh. + +End STMT_FUNC. diff --git a/src/lib/ArithInv.v b/src/lib/ArithInv.v new file mode 100755 index 0000000..5805baf --- /dev/null +++ b/src/lib/ArithInv.v @@ -0,0 +1,108 @@ +Require Import ZArith. +Require Import cclib.Coqlib. +Require Import cclib.Integers. + +Lemma Zgt_bool_false: forall n m : Z, (n >? m)%Z = false -> (n <= m)%Z. +Proof. + intros. + assert (H1 := Zgt_cases n m). + rewrite H in H1. + auto. +Qed. + +Lemma Zlt_bool_false: forall n m : Z, (n (n >= m)%Z. +Proof. + intros. + assert (H1 := Zlt_cases n m). + rewrite H in H1. + auto. +Qed. + +Lemma zle_bool_false: forall n m : Z, (proj_sumbool (zle n m)) = false -> (n > m)%Z. +Proof. + intros. + destruct (zle n m); [discriminate | assumption ]. +Qed. + +Lemma zle_bool_true: forall n m : Z, (proj_sumbool (zle n m)) = true -> (n <= m)%Z. +Proof. + intros. + intros. + destruct (zle n m); [assumption| discriminate ]. +Qed. + +Lemma zlt_bool_false: forall n m : Z, (proj_sumbool (zlt n m)) = false -> (n >= m)%Z. +Proof. + intros. + destruct (zlt n m); [discriminate | assumption ]. +Qed. + +Lemma zlt_bool_true: forall n m : Z, (proj_sumbool (zlt n m)) = true -> (n < m)%Z. +Proof. + intros. + intros. + destruct (zlt n m); [assumption| discriminate ]. +Qed. + +Lemma Int256eq_true : forall x y, Int256.eq x y = true -> x = y. +Proof. + intros. + assert (FOO := Int256.eq_spec x y). + rewrite H in FOO. + assumption. +Qed. + +Lemma Int256eq_false : forall x y, Int256.eq x y = false -> x <> y. +Proof. + intros. + assert (FOO := Int256.eq_spec x y). + rewrite H in FOO. + assumption. +Qed. + +Lemma geb_ge: forall n m : Z, (n >=? m) = true -> n >= m. +Proof. + intros. + rewrite Z.geb_le in H. omega. +Qed. + +Lemma ngeb_lt: forall n m : Z, (n >=? m) = false -> n < m. +Proof. + intros. + assert (H2:= Zge_cases n m). + rewrite H in H2. + auto. +Qed. + +Lemma if_and : forall (a b : bool), + (if a then b else false) = true -> a = true /\ b = true. +Proof. + destruct a, b; simpl; intros; auto. +Qed. + + +Ltac inv_arith := + repeat match goal with + | [H : negb _ = true |- _ ] => rewrite Bool.negb_true_iff in H + | [H : negb _ = false |- _ ] => rewrite <- Bool.negb_false_iff in H + | [H : (andb _ _) = true |- _ ] => apply andb_prop in H; destruct H + | [H : (if _ then _ else false) = true |- _] => apply if_and in H; destruct H + + | [H : (Int256.eq _ _) = true |- _] => apply Int256eq_true in H + | [H : (Int256.eq _ _) = false |- _] => apply Int256eq_false in H + + | [H : (_ >? _)%Z = true |- _] => rewrite <- Zgt_is_gt_bool in H + | [H : (_ >? _)%Z = false |- _] => apply Zgt_bool_false in H + | [H : (_ rewrite <- Zlt_is_lt_bool in H + | [H : (_ apply Zlt_bool_false in H + | [H : (_ =? _)%Z = true |- _] => rewrite Z.eqb_eq in H + | [H : (_ =? _)%Z = false |- _] => rewrite Z.eqb_neq in H + | [H : (_ >=? _)%Z = true |- _] => apply geb_ge in H + | [H : (_ >=? _)%Z = false |- _] => apply ngeb_lt in H + + | [H : (proj_sumbool (zle _ _)) = true |- _] => apply zle_bool_true in H + | [H : (proj_sumbool (zle _ _)) = false |- _] => apply zle_bool_false in H + | [H : (proj_sumbool (zlt _ _)) = true |- _] => apply zlt_bool_true in H + | [H : (proj_sumbool (zlt _ _)) = false |- _] => apply zlt_bool_false in H + end. + diff --git a/src/lib/IndexedMaps.v b/src/lib/IndexedMaps.v new file mode 100755 index 0000000..87cd225 --- /dev/null +++ b/src/lib/IndexedMaps.v @@ -0,0 +1,651 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Applicative finite maps are the main data structure used in this + project. A finite map associates data to keys. The two main operations + are [set k d m], which returns a map identical to [m] except that [d] + is associated to [k], and [get k m] which returns the data associated + to key [k] in map [m]. In this library, we distinguish two kinds of maps: +- Trees: the [get] operation returns an option type, either [None] + if no data is associated to the key, or [Some d] otherwise. +- Maps: the [get] operation always returns a data. If no data was explicitly + associated with the key, a default data provided at map initialization time + is returned. + + In this library, we provide efficient implementations of trees and + maps whose keys range over the type [positive] of binary positive + integers or any type that can be injected into [positive]. The + implementation is based on radix-2 search trees (uncompressed + Patricia trees) and guarantees logarithmic-time operations. An + inefficient implementation of maps as functions is also provided. +*) + +Require Export Coq.Program.Equality. + +Require Import Coqlib. +Require Import Maps. + +Record Isomorphism {A} (X Y : A) : Type := { + iso_f : forall F : A -> Type, F X -> F Y; + iso_g : forall F : A -> Type, F Y -> F X + (* := fun F => iso_f (fun x => F x -> F X) id *) +}. + +Arguments iso_f {_ _ _} _ [_] _. +Arguments iso_g {_ _ _} _ [_] _. +Coercion iso_f : Isomorphism >-> Funclass. + +Notation "X ~~~ Y" := (Isomorphism X Y) (at level 70, no associativity). + +Definition identity_automorphism {A X} : @Isomorphism A X X + := {| iso_f F fx := fx; iso_g f fx := fx |}. + +Definition iso_concat {A}{X Y Z : A}(iso_xy : X ~~~ Y)(iso_yz : Y ~~~ Z) := + {| iso_f F fx := iso_yz F (iso_xy F fx); + iso_g F fy := iso_g iso_xy (iso_g iso_yz fy) |}. + +Definition iso_inverse {A}{X Y : A}(iso : X ~~~ Y) := + {| iso_f := iso_g iso; iso_g := iso_f iso |}. + +Notation "p @ q" := (iso_concat p q) (at level 40, left associativity). +Notation "! p" := (iso_inverse p) (at level 35, right associativity). + +Definition isomorphism_compose {A}{X Y Z : A}(iso_xy : X ~~~ Y) + (iso_yz : Y ~~~ Z) : + forall F fx, iso_yz F (iso_xy F fx) = (iso_xy @ iso_yz) F fx. +Proof. + intros F fx; reflexivity. +Qed. + +Lemma isomorphism_decidable_neq_absurd {A}{X Y : A}(dec : A -> bool) : + X ~~~ Y -> dec X = true -> dec Y = false -> False. +Proof. + intros iso dec_x dec_y. + set (F := fun Z : A => match dec Z with + | true => True + | false => False + end). + assert (fx : F X). + { unfold F; rewrite dec_x; exact I. } + + assert (fy : F Y). + { apply iso, fx. } + + unfold F in fy. + rewrite dec_y in fy. + exact fy. +Qed. + +(** Defining the isomorphism using Leibniz equality means that we have to + rely on some properties of functors. One of them, especially, is based + on paramatricity, which can't be proved inside Coq, and is axiomatized + as [functor_identity_paramatricity]. + + Hopefully it is consistent. It implies eq_rect_eq, which is already + implied by proof irrelevance required by CompCert. +*) +Axiom functor_identity_paramatricity : + forall (A : Type)(X : A)(f : forall F, F X -> F X) + (F : A -> Type)(x : F X), x = f F x. + +Module AxiomImplication. + Lemma eq_rect_eq : + forall (U : Type) (p : U) (Q : U -> Type) (x : Q p) (h : p = p), + x = eq_rect p Q x p h. + Proof. + intros. + exact (functor_identity_paramatricity U p + (fun Q x => eq_rect p Q x p h) Q x). + Qed. +End AxiomImplication. + +Lemma isomorphism_identity_gf' {A}{X Y : A}(iso1 iso2 : X ~~~ Y)(F : A -> Type) + (x : F X) : x = iso_g iso1 (iso_f iso2 x). +Proof. + set (iso_comp := iso2 @ !iso1). + change (iso_g iso1 (iso_f iso2 x)) with (iso_f iso_comp x). + apply functor_identity_paramatricity. +Qed. + +Lemma isomorphism_identity_fg' {A}{X Y : A}(iso1 iso2 : X ~~~ Y)(F : A -> Type) + (y : F Y) : y = iso_f iso1 (iso_g iso2 y). +Proof. + set (iso_comp := !iso1 @ iso2). + change (iso_f iso1 (iso_g iso2 y)) with (iso_g iso_comp y). + apply functor_identity_paramatricity. +Qed. + +Lemma isomorphism_identity_gf {A}{X Y : A}(iso : X ~~~ Y)(F : A -> Type) + (x : F X) : x = iso_g iso (iso_f iso x). +Proof (isomorphism_identity_gf' iso iso F x). + +Lemma isomorphism_identity_fg {A}{X Y : A}(iso : X ~~~ Y)(F : A -> Type) + (y : F Y) : y = iso_f iso (iso_g iso y). +Proof (isomorphism_identity_fg' iso iso F y). + +Lemma isomorphism_unique {A}{X Y : A}(iso1 iso2 : X ~~~ Y)(F : A -> Type) + (x : F X) : iso_f iso1 x = iso_f iso2 x. +Proof. + transitivity (iso_f iso2 (iso_g iso2 (iso_f iso1 x))). + apply isomorphism_identity_fg. + f_equal; symmetry. + apply isomorphism_identity_gf'. +Qed. + +Lemma isomorphism_injective {A}{X Y : A}(iso : X ~~~ Y) : + forall F x1 x2, iso F x1 = iso F x2 -> x1 = x2. +Proof. + intros F x1 x2 eq. + rewrite (isomorphism_identity_gf iso F x1). + rewrite eq. + symmetry; apply isomorphism_identity_gf. +Qed. + +(* To avoid useless definitions of inductors in extracted code. *) +Local Unset Elimination Schemes. +Local Unset Case Analysis Schemes. + +(** * An implementation of indexed trees over type [positive] + Element types are indexed by a tree [PTree.t A] along with a projection + [proj] that extract the actual [Type] from [A]. *) +Module Type TypeProjection. + Variable A : Type. + Variable proj : A -> Type. +End TypeProjection. + +Module IPTree(TP : TypeProjection). + Definition elt := positive. + Definition elt_eq := peq. + + Definition OptionalType (a : option TP.A) : Type := + match a with + | None => unit + | Some a' => TP.proj a' + end. + + Inductive tree : PTree.t TP.A -> Type := + | Leaf : tree PTree.Leaf + | Node : forall l r o, tree l -> OptionalType o -> tree r -> + tree (PTree.Node l o r). + + Arguments Node [l r]. + Scheme tree_ind := Induction for tree Sort Prop. + + Definition get_index {idx}(_ : tree idx) := idx. + + Definition t := tree. + + Definition empty : t (PTree.empty TP.A) := Leaf. + + Fixpoint get {idx}(i : positive)(m : t idx){struct i} : + OptionalType (PTree.get i idx). + Proof. + destruct idx. + - destruct i; simpl; exact tt. + - set (tuple := match m in (tree t') + return match t' with + | PTree.Leaf => t PTree.Leaf * unit * t PTree.Leaf + | PTree.Node l o r => t l * OptionalType o * t r + end with + | Leaf => ((Leaf, tt), Leaf) + | Node _ _ _ m1 o0 m2 => ((m1, o0), m2) + end). + destruct i; simpl; try apply get; apply tuple. + Defined. + + Definition get_eq {idx}(i : positive){OA}(m : t idx) : + PTree.get i idx ~~~ OA -> OptionalType OA + := fun iso => iso OptionalType (get i m). + + Fixpoint set {idx}(i : positive)(A : TP.A)(v : TP.proj A)(m : t idx){struct i} : + t (PTree.set i A idx). + Proof. + destruct idx. + - destruct i; simpl; apply Node; try apply (set _ i); + solve [ exact Leaf | exact tt | exact v ]. + - set (tuple := match m in (tree t') + return match t' with + | PTree.Leaf => t PTree.Leaf * unit * t PTree.Leaf + | PTree.Node l o r => t l * OptionalType o * t r + end with + | Leaf => ((Leaf, tt), Leaf) + | Node _ _ _ m1 o0 m2 => ((m1, o0), m2) + end). + destruct i; simpl; apply Node; try apply (set _ i); + solve [ apply tuple | exact v ]. + Defined. + + Fixpoint remove {idx}(i : positive)(m : t idx){struct i} : + t (PTree.remove i idx). + Proof. + (* Comments in between tactics roughly match cases of PTree.remove *) + + destruct i; destruct idx. + (* match i, m with *) + + (* | xI i, Leaf *) + - exact Leaf. + - set (tuple := match m in (tree t') + return match t' with + | PTree.Leaf => t PTree.Leaf * unit * t PTree.Leaf + | PTree.Node l o r => t l * OptionalType o * t r + end with + | Leaf => (Leaf, tt, Leaf) + | Node _ _ _ m1 o0 m2 => (m1, o0, m2) + end). + set (m1 := fst (fst tuple)). set (o0 := snd (fst tuple)). + set (m2 := snd tuple). + simpl; destruct idx1. destruct o. + (* | xI i, Node (Some a) Leaf o0 m2 *) + + exact (Node _ m1 o0 (remove _ i m2)). + (* | xI i, Node None Leaf _ m2 *) + + set (rm_i_r := PTree.remove i idx2). + (* match PTree.remove i r with *) + refine (match rm_i_r as t' + return (PTree.remove i idx2 ~~~ t' -> + t match t' with + | PTree.Leaf => PTree.Leaf + | PTree.Node l o r => PTree.Node PTree.Leaf None + (PTree.Node l o r) + end) with + (* | PTree.Leaf *) + | PTree.Leaf => fun _ => Leaf + | PTree.Node l' o' r' => fun rm_iso => _ (* !!! *) + end identity_automorphism). + (* | mm *) + refine (Node None Leaf tt _). + apply rm_iso. + exact (remove _ i m2). + (* end *) + (* | xI i, Node o m1 o0 m2 *) + + exact (Node _ m1 o0 (remove _ i m2)). + + (* | xO i, Leaf *) + - exact Leaf. + - set (tuple := match m in (tree t') + return match t' with + | PTree.Leaf => t PTree.Leaf * unit * t PTree.Leaf + | PTree.Node l o r => t l * OptionalType o * t r + end with + | Leaf => (Leaf, tt, Leaf) + | Node _ _ _ m1 o0 m2 => (m1, o0, m2) + end). + set (m1 := fst (fst tuple)). set (o0 := snd (fst tuple)). + set (m2 := snd tuple). + simpl; destruct o. 2: destruct idx2. + (* | xO i, Node (Some a) m1 o0 m2 *) + + exact (Node _ (remove _ i m1) o0 m2). + (* | xO i, Node None m1 _ m2 *) + + set (rm_i_l := PTree.remove i idx1). + (* match PTree.remove i l with *) + refine (match rm_i_l as t' + return (PTree.remove i idx1 ~~~ t' -> + t match t' with + | PTree.Leaf => PTree.Leaf + | PTree.Node l o r => PTree.Node (PTree.Node l o r) + None PTree.Leaf + end) with + (* | PTree.Leaf *) + | PTree.Leaf => fun _ => Leaf + | PTree.Node l' o' r' => fun rm_iso => _ (* !!! *) + end identity_automorphism). + (* | mm *) + refine (Node None _ tt m2). + apply rm_iso. + exact (remove _ i m1). + (* end *) + (* | xO i, Node None m1 o0 m2 *) + + exact (Node None (remove _ i m1) tt m2). + + (* | xH, Leaf *) + - exact Leaf. + - set (tuple := match m in (tree t') + return match t' with + | PTree.Leaf => t PTree.Leaf * unit * t PTree.Leaf + | PTree.Node l o r => t l * OptionalType o * t r + end with + | Leaf => (Leaf, tt, Leaf) + | Node _ _ _ m1 o0 m2 => (m1, o0, m2) + end). + set (m1 := fst (fst tuple)). set (o0 := snd (fst tuple)). + set (m2 := snd tuple). + simpl; destruct idx1. destruct idx2. + (* | xH, Node o Leaf o0 Leaf *) + + exact Leaf. + (* | xH, Node o Leaf o0 m2 *) + + exact (Node None Leaf tt m2). + (* | xH, Node o m1 o0 m2 *) + + exact (Node None m1 tt m2). + (* end *) + Defined. + + Fixpoint remove_eq {idx}(i : positive) A (m : t (PTree.set i A idx)){struct i} : + PTree.get i idx ~~~ None -> t idx. + Proof. + destruct i; simpl in m; unfold PTree.get; try fold (@PTree.get TP.A i); + destruct idx; try (intros _; exact Leaf). + + (* xI i *) + - intros idx2_i_iso_none. + set (tuple := match m in (tree t') + return match t' with + | PTree.Leaf => t PTree.Leaf * unit * t PTree.Leaf + | PTree.Node l o r => t l * OptionalType o * t r + end with + | Leaf => (Leaf, tt, Leaf) + | Node _ _ _ m1 o0 m2 => (m1, o0, m2) + end). + set (m1 := fst (fst tuple)). set (o0 := snd (fst tuple)). + set (m2 := snd tuple). + refine (Node _ m1 o0 _). + apply remove_eq with i A; assumption. + (* xO i *) + - intros idx1_i_iso_none. + set (tuple := match m in (tree t') + return match t' with + | PTree.Leaf => t PTree.Leaf * unit * t PTree.Leaf + | PTree.Node l o r => t l * OptionalType o * t r + end with + | Leaf => (Leaf, tt, Leaf) + | Node _ _ _ m1 o0 m2 => (m1, o0, m2) + end). + set (m1 := fst (fst tuple)). set (o0 := snd (fst tuple)). + set (m2 := snd tuple). + refine (Node _ _ o0 m2). + apply remove_eq with i A; assumption. + (* xH *) + - intros o_iso_none. + set (tuple := match m in (tree t') + return match t' with + | PTree.Leaf => t PTree.Leaf * unit * t PTree.Leaf + | PTree.Node l o r => t l * OptionalType o * t r + end with + | Leaf => (Leaf, tt, Leaf) + | Node _ _ _ m1 o0 m2 => (m1, o0, m2) + end). + set (m1 := fst (fst tuple)). set (m2 := snd tuple). + refine (Node _ m1 _ m2). + apply o_iso_none, tt. + Defined. + Global Arguments remove_eq {_} _ {_} _ _. + + Theorem gempty (i : positive)(iso : PTree.Leaf ! i ~~~ None) : + iso OptionalType (get i empty) = tt. + Proof. + induction i; symmetry; + apply (functor_identity_paramatricity _ _ (iso_f iso) OptionalType). + Qed. + + Lemma gleaf (i : positive)(iso : PTree.Leaf ! i ~~~ None) : + iso OptionalType (get i Leaf) = tt. + Proof (gempty i iso). + + Theorem get_eq_fusion {idx OA OA'}(i : positive)(m : t idx) + (iso : PTree.get i idx ~~~ OA)(iso' : OA ~~~ OA') : + iso' OptionalType (get_eq i m iso) = get_eq i m (iso @ iso'). + Proof. + unfold get_eq. + rewrite isomorphism_compose. + reflexivity. + Qed. + + Theorem gss {idx A}(i : positive)(x : TP.proj A)(m : t idx) + (iso : (PTree.set i A idx) ! i ~~~ Some A) : + get i (set i A x m) = iso_g iso (F := OptionalType) x. + Proof. + generalize dependent idx; induction i. + - destruct m; simpl; auto. + - destruct m; simpl; auto. + - destruct m; simpl; intro iso; + apply (functor_identity_paramatricity _ _ (iso_g iso) OptionalType). + Qed. + + Theorem gess' {idx A OA}(i : positive)(x : TP.proj A)(m : t idx) + iso (iso' : Some A ~~~ OA) : + get_eq i (OA := OA) (set i A x m) iso = iso' OptionalType x. + Proof. + unfold get_eq. + set (iso_comp := iso @ !iso'). + rewrite (gss _ _ _ iso_comp). + symmetry. + apply (isomorphism_identity_fg iso _ (iso' OptionalType x)). + Qed. + + Theorem gess {idx A}(i : positive)(x : TP.proj A)(m : t idx) iso : + get_eq i (OA := Some A) (set i A x m) iso = x. + Proof (gess' i x m iso identity_automorphism). + + Definition gleaf_iso {i} : (@PTree.Leaf TP.A) ! i ~~~ None. + Proof. destruct i; exact identity_automorphism. Defined. + + Theorem gso {idx A}(i j : positive)(x : TP.proj A)(m : t idx) + iso : i <> j -> get i (set j A x m) = iso_f iso (get i m). + Proof. + generalize dependent idx; generalize dependent j. + induction i. + - intros j idx m iso neq. + destruct idx. + + destruct j; simpl in * |- *. + * rewrite <- (gleaf i gleaf_iso). + set (iso_comp := (@gleaf_iso i) @ iso). + change (iso_f iso (iso_f gleaf_iso (get i Leaf))) with + (iso_f iso_comp (get i Leaf)). + apply IHi; congruence. + * destruct i; apply functor_identity_paramatricity. + * destruct i; apply functor_identity_paramatricity. + + destruct j; simpl in * |- *. + * apply IHi; congruence. + * apply functor_identity_paramatricity. + * apply functor_identity_paramatricity. + - intros j idx m iso neq. + destruct m. + + destruct j; simpl in * |- *. + * destruct i; apply functor_identity_paramatricity. + * rewrite <- (gleaf i gleaf_iso). + set (iso_comp := (@gleaf_iso i) @ iso). + change (iso_f iso (iso_f gleaf_iso (get i Leaf))) with + (iso_f iso_comp (get i Leaf)). + apply IHi; congruence. + * destruct i; apply functor_identity_paramatricity. + + destruct j; simpl in * |- *. + * apply functor_identity_paramatricity. + * apply IHi; congruence. + * apply functor_identity_paramatricity. + - intros j idx m iso neq. + destruct m. + + destruct j; simpl in * |- *. + * apply (functor_identity_paramatricity _ None _ OptionalType). + * apply (functor_identity_paramatricity _ None _ OptionalType). + * congruence. + + destruct j; simpl in * |- *. + * apply functor_identity_paramatricity. + * apply functor_identity_paramatricity. + * congruence. + Qed. + + Theorem geso {idx A OA}(i j : positive)(x : TP.proj A)(m : t idx) iso iso' : + i <> j -> get_eq i (set j A x m) iso = @get_eq _ i OA m iso'. + Proof. + unfold get_eq. + intros neq. + rewrite (gso _ _ x m (iso' @ !iso) neq). + set (iso_comp := (iso' @ !iso) @ iso). + change (iso_f iso (iso_f (iso' @ !iso) (get i m))) with + (iso_f iso_comp (get i m)). + apply isomorphism_unique. + Qed. + + Theorem gres {idx A}(i : positive)(m : t (PTree.set i A idx)) iso : + get i (remove_eq i m iso) = iso_g iso (F := OptionalType) tt. + Proof. + generalize dependent idx; induction i; destruct idx. + (* xI i, PTree.Leaf *) + - intros m iso. + apply (functor_identity_paramatricity _ _ (iso_g iso) OptionalType). + (* xI i, PTree.Node *) + - intros m; apply IHi. + (* xO i, PTree.Leaf *) + - intros m iso. + apply (functor_identity_paramatricity _ _ (iso_g iso) OptionalType). + (* xO i, PTree.Node *) + - intros m; apply IHi. + (* xH, PTree.Leaf *) + - intros m iso. + apply (functor_identity_paramatricity _ _ (iso_g iso) OptionalType). + (* xH, PTree.None *) + - intros m iso. + reflexivity. + Qed. + + Theorem geres' {idx A OA}(i : positive)(m : t (PTree.set i A idx)) iso iso' : + get_eq i (OA := OA) (remove_eq i m iso) iso' = iso_g (!iso' @ iso) tt. + Proof. + unfold get_eq, iso_concat, iso_inverse, iso_g at 1, iso_g at 1. + f_equal. + apply gres. + Qed. + + Theorem geres {idx A}(i : positive)(m : t (PTree.set i A idx)) iso : + get_eq i (OA := None) (remove_eq i m iso) iso = tt. + Proof. + rewrite (isomorphism_identity_fg iso OptionalType tt). + apply geres'. + Qed. + + Lemma rleaf {A}(i : positive)(m : t (PTree.set i A PTree.Leaf)) iso : + remove_eq i m iso = Leaf. + Proof. destruct i; reflexivity. Qed. + + Theorem greo {idx A}(i j : positive)(m : t (PTree.set j A idx)) iso iso' : + i <> j -> get i (remove_eq j m iso) = iso_f iso' (get i m). + Proof. + generalize dependent idx; generalize dependent j. + induction i; destruct idx. + (* xI i, PTree.Leaf *) + - intros m iso iso' neq. + destruct j; simpl. + + set (m2 := let (_, H) := + match m in (tree t') + return match t' with + | PTree.Leaf => t PTree.Leaf * unit * t PTree.Leaf + | PTree.Node l o r => t l * OptionalType o * t r + end with + | Leaf => (Leaf, tt, Leaf) + | Node l r o1 m1 o0 m2 => (m1, o0, m2) + end + in H). + assert (gleaf_iso OptionalType (get i (remove_eq j m2 gleaf_iso)) = tt). + destruct i; destruct j; simpl; reflexivity. + rewrite <- H at 1. + apply (isomorphism_injective (iso_inverse (@gleaf_iso i)) _ + (gleaf_iso OptionalType (get i (remove_eq j m2 gleaf_iso)))). + (* iso_f (!gleaf_iso) (iso_f gleaf_iso (get i ...)) => get i ... *) + unfold iso_inverse at 1, iso_f at 1; + rewrite <- (isomorphism_identity_gf (@gleaf_iso i)). + rewrite isomorphism_compose. + apply IHi; congruence. + + destruct i; apply (functor_identity_paramatricity _ _ iso' OptionalType). + + destruct i; apply (functor_identity_paramatricity _ _ iso' OptionalType). + (* xI i, PTree.Node *) + - intros m iso iso' neq. + destruct j; simpl. + + apply IHi; congruence. + + apply functor_identity_paramatricity. + + apply functor_identity_paramatricity. + (* xO i, PTree.Leaf *) + - intros m iso iso' neq. + destruct j; simpl. + + destruct i; apply (functor_identity_paramatricity _ _ iso' OptionalType). + + set (m1 := let (H, _) := let (H, _) := + match m in (tree t') + return match t' with + | PTree.Leaf => t PTree.Leaf * unit * t PTree.Leaf + | PTree.Node l o r => t l * OptionalType o * t r + end with + | Leaf => (Leaf, tt, Leaf) + | Node l r o1 m1 o0 m2 => (m1, o0, m2) + end + in H in H). + assert (gleaf_iso OptionalType (get i (remove_eq j m1 gleaf_iso)) = tt). + destruct i; destruct j; simpl; reflexivity. + rewrite <- H at 1. + apply (isomorphism_injective (iso_inverse (@gleaf_iso i)) _ + (gleaf_iso OptionalType (get i (remove_eq j m1 gleaf_iso)))). + (* iso_f (!gleaf_iso) (iso_f gleaf_iso (get i ...)) => get i ... *) + unfold iso_inverse at 1, iso_f at 1; + rewrite <- (isomorphism_identity_gf (@gleaf_iso i)). + rewrite isomorphism_compose. + apply IHi; congruence. + + destruct i; apply (functor_identity_paramatricity _ _ iso' OptionalType). + (* xO i, PTree.Node *) + - intros m iso iso' neq. + destruct j; simpl. + + apply functor_identity_paramatricity. + + apply IHi; congruence. + + apply functor_identity_paramatricity. + (* xH, PTree.Leaf *) + - intros m iso iso' neq. + destruct j; try congruence; simpl in m |- *; + dependent destruction m; destruct o0; + apply (functor_identity_paramatricity _ _ (iso_f iso') OptionalType). + (* xH, PTree.Node *) + - intros m iso iso' neq. + destruct j; simpl. + + apply (functor_identity_paramatricity _ _ (iso_f iso') OptionalType). + + apply (functor_identity_paramatricity _ _ (iso_f iso') OptionalType). + + congruence. + (* TODO: solve universe inconsistency *) + Admitted. + + Theorem gereo {idx A OA}(i j : positive)(m : t (PTree.set j A idx)) + iso iso' iso'' : + i <> j -> get_eq i (OA := OA) (remove_eq j m iso) iso' + = iso_f iso' (get_eq i m iso''). + Proof. + intros neq; unfold get_eq. + f_equal; apply greo; assumption. + Qed. + +End IPTree. + +Module PTreeAux. + Theorem smap1: + forall (A B: Type) (f: A -> B) (a: A) (i: PTree.elt) (m: PTree.t A), + PTree.set i (f a) (PTree.map1 f m) = PTree.map1 f (PTree.set i a m). + Proof. + induction i; destruct m; simpl; try rewrite <- IHi; reflexivity. + Qed. +End PTreeAux. + +Module PMapAux. + Theorem smap: + forall (A B: Type) (f: A -> B) (a: A) (i: PMap.elt) (m: PMap.t A), + PMap.set i (f a) (PMap.map f m) = PMap.map f (PMap.set i a m). + Proof. + intros. + unfold PMap.map, PMap.set; simpl. + f_equal; apply PTreeAux.smap1. + Qed. +End PMapAux. + +Module ZMapAux. + Theorem smap: + forall (A B: Type) (f: A -> B) (a: A) (i: ZMap.elt) (m: ZMap.t A), + ZMap.set i (f a) (ZMap.map f m) = ZMap.map f (ZMap.set i a m). + Proof. + intros. + apply PMapAux.smap. + Qed. +End ZMapAux. diff --git a/src/lib/LangDef.v b/src/lib/LangDef.v new file mode 100755 index 0000000..8ec3657 --- /dev/null +++ b/src/lib/LangDef.v @@ -0,0 +1,125 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(* Standard library modules *) +Require Import BinInt. +Require Omega. + +(* CompCert modules *) +(*Require Import compcert.common.Events.*) +Require Import cclib.Coqlib. +(*Require Import compcert.cfrontend.Clight. +Require Import compcert.cfrontend.ClightBigstep. *) + +(* CompCertX modules *) +(* Require compcertx.common.EventsX. *) + +Local Open Scope Z_scope. + + +(* TODO: figure out a better statement of these, and move them to the backend directory. *) + +(* +Definition exec_stmt `{compiler_config_ops : CompilerConfigOps} + m0 ge env le (m : mem) s t le' m' out + := let wbwim := EventsX.writable_block_with_init_mem_ops in + let wb := Events.writable_block_with_init_mem_writable_block_ops m0 in + ClightBigstep.exec_stmt ge (fun _ => function_entry2) env le m s t le' m' out. + +Definition eval_funcall {mem}`{compiler_config_ops : CompilerConfigOps mem} + m0 ge (m : mem) f vargs t m' vres + := let wbwim := EventsX.writable_block_with_init_mem_ops in + let wb := Events.writable_block_with_init_mem_writable_block_ops m0 in + ClightBigstep.eval_funcall ge (fun _ => Clight.function_entry2) + m f vargs t m' vres. +*) + +Lemma min_ex: forall (P: Z -> Prop) lo hi, + (forall n, {P n} + {~ P n}) -> + {n : Z | lo <= n < hi /\ P n /\ forall n', lo <= n' < n -> ~ P n'} + + {forall n, lo <= n < hi -> ~ P n}. +Proof. + intros. + assert(HP: forall k, 0 <= k -> + {n : Z | lo <= n < lo + k /\ P n /\ forall n' : Z, lo <= n' < n -> ~ P n'} + + {forall n : Z, lo <= n < lo + k -> ~ P n}). + apply natlike_rec2. + right. + intros. + omega. + + intros z HR HT. + destruct HT. + left. + destruct s as [n[HT HM]]. + exists n. + split; trivial. + omega. + + specialize (H (lo + z)). + destruct H. + left. + exists (lo + z). + split; auto. + omega. + + right. + intros. + destruct (zeq n1 (lo + z)). + subst. + trivial. + apply n. + omega. + + destruct (zlt lo hi). + replace hi with (lo + (hi - lo)) by omega. apply HP. omega. + right; intros. omegaContradiction. + +Qed. + +Lemma min_ex_found_same: + forall (P: Z -> Prop) lo hi n1 n2, + (lo <= n1 < hi /\ P n1 /\ forall n', lo <= n' < n1 -> ~ P n') -> + (lo <= n2 < hi /\ P n2 /\ forall n', lo <= n' < n2 -> ~ P n') -> + n1 = n2. +Proof. + intros ? ? ? ? ? (n1_range & Pn1 & n1_first) (n2_range & Pn2 & n2_first). + destruct (Z.lt_total n1 n2) as [ ne | [ n_eq | ne ]]; try exact n_eq. + - exfalso. + refine (n2_first _ _ Pn1). + split; [ apply n1_range | apply ne ]. + - exfalso. + refine (n1_first _ _ Pn2). + split; [ apply n2_range | apply ne ]. +Qed. + +Lemma min_ex_found_not_found_contradict: + forall (P: Z -> Prop) lo hi n, + (lo <= n < hi /\ P n /\ forall n', lo <= n' < n -> ~ P n') -> + (forall n, lo <= n < hi -> ~ P n) -> + False. +Proof. + intros ? ? ? ? (n_range & Pn & _) not_found. + eapply not_found; eassumption. +Qed. diff --git a/src/lib/Lens.v b/src/lib/Lens.v new file mode 100755 index 0000000..b43ab8a --- /dev/null +++ b/src/lib/Lens.v @@ -0,0 +1,555 @@ +Require Import Coq.Program.Basics. +Require Import Coq.Classes.Morphisms. +Require Import Coq.Classes.RelationClasses. +Require Export Coq.Classes.RelationPairs. +Require Import Coq.Setoids.Setoid. + +(** * Interface *) + +(** Lenses are identified by the corresponding projection, which is + also used to define the [get] operation. Note that because we use + the letter [S] to denote the source type, which conflicts with the + successor function on [nat] from the standard library, we need to + use a verbose style below to avoid arguments being named [S0] or + something such. (This is an obvious software engineering quirk + with the way Coq generates identifier.) *) + +Class LensOps {S V: Type} (π: S -> V) := { + set: V -> S -> S +}. + +Arguments set {S V} π {_} _ _ : simpl never. + +Class LensSetGet {S V} π + `{lens_ops: LensOps S V π} := +{ + lens_set_get s: + set π (π s) s = s +}. + +Class LensGetSet {S V} π + `{lens_ops: LensOps S V π} := +{ + lens_get_set v s: + π (set π v s) = v +}. + +Class LensSetSet {S V} π + `{lens_set: LensOps S V π} := +{ + lens_set_set u v s: + set π u (set π v s) = set π u s +}. + +Class Lens {S V} π `{lens_ops: LensOps S V π} := { + lens_lens_set_get :> LensSetGet π; + lens_lens_get_set :> LensGetSet π; + lens_lens_set_set :> LensSetSet π +}. + + +(** * Theory *) + +(** ** Getters are measures, cf. [Coq.Classes.RelationPairs] *) + +Instance lens_get_measure {S V} `{LensOps S V}: + Measure (π). + +(** ** The [modify] operation *) + +Section MODIFY. + Context {S V π} `{HSV: Lens S V π}. + + Definition modify (f: V -> V) (s: S) := + set π (f (π s)) s. + + Lemma lens_unfold_modify f s: + modify f s = set π (f (π s)) s. + Proof. + reflexivity. + Qed. +End MODIFY. + +Arguments modify {S V} π {_} _ _. + +(** ** The [same_context] relation *) + +Section SAMECONTEXT. + Context {S V π} `{lens_ops: LensOps S V π} `{HSV: LensSetSet S V π}. + + Definition same_context: relation S := + fun s t => forall v, set π v s = set π v t. + + Lemma lens_set_same_context v s: + same_context (set π v s) s. + Proof. + intro u. + apply lens_set_set. + Qed. + + Lemma lens_modify_same_context f s: + same_context (modify π f s) s. + Proof. + intro u. + unfold modify. + apply lens_set_set. + Qed. + + Global Instance same_context_equiv: Equivalence same_context. + Proof. + split. + * intros s v. + reflexivity. + * intros s t Hst u. + symmetry; now auto. + * intros s1 s2 s3 H12 H23 u. + transitivity (set π u s2); now auto. + Qed. + + Global Instance same_context_set_mor: + Proper (eq ==> same_context ==> eq) (set π). + Proof. + intros u v Huv s t Hst. + subst. + apply Hst. + Qed. + + Global Instance same_context_modify_mor f: + Proper (same_context ==> same_context) (modify π f). + Proof. + intros s t Hst u. + unfold modify. + rewrite !lens_set_set. + apply Hst. + Qed. +End SAMECONTEXT. + +Arguments same_context {S V} π {_} _ _. + +(** ** Consequences of [LensGetSet] *) + +Section GETSET. + Context {S V} `{Hgs: LensGetSet S V}. + + Lemma lens_get_modify f s: + π (modify π f s) = f (π s). + Proof. + apply lens_get_set. + Qed. + + Theorem lens_eq_set (s: S) (u v: V): + set π u s = set π v s <-> u = v. + Proof. + split; intros. + * rewrite <- (lens_get_set u s). + rewrite <- (lens_get_set v s). + apply f_equal. + assumption. + * apply (f_equal (fun x => set π x s)). + assumption. + Qed. +End GETSET. + +(** ** Consequences of [LensSetGet] *) + +Section SETGET. + Context {S V} `{HSV: LensSetGet S V}. + + (* same_context ∧ (eq @@ π) ⇔ eq *) + Lemma lens_same_context_eq s1 s2: + same_context π s1 s2 -> + π s1 = π s2 -> + s1 = s2. + Proof. + intros Hc Hv. + rewrite <- (lens_set_get s1). + rewrite <- (lens_set_get s2). + rewrite Hv. + apply Hc. + Qed. +End SETGET. + +(** ** Orthogonal lenses *) + +(** We say that two lenses on [S] are orthogonal when they give us + access to two independent parts of the larger structure [S], + ie. modifying one will not affect the other. The property + [lens_ortho_setr_setl] below is sufficient to express this. + + Note that for any two orthogonal lenses [π] and [ρ], you should + never declare instances of both [OrthogonalLenses π ρ] and + [OrthogonalLenses ρ π], since that would produce a loop in the + [autorewrite] rules, which use [ortho_lenses_set_set]. *) + +Class Orthogonal {S U V} (π: S -> U) (ρ: S -> V) + `{πs: !LensOps π} + `{ρs: !LensOps ρ}: Prop := +{ + lens_ortho_setr_setl u v s: + set ρ v (set π u s) = set π u (set ρ v s) +}. + +Section ORTHOGONAL_LENSES. + Context {S U V} (π: S -> U) (ρ: S -> V). + Context `{πops: !LensOps π}. + Context `{ρops: !LensOps ρ}. + Context `{Hπρ: !Orthogonal π ρ}. + + Lemma lens_ortho_getl_setr `{Hπgs: !LensGetSet π} `{Hπsg: !LensSetGet π} s v: + π (set ρ v s) = π s. + Proof. + rewrite <- (lens_set_get (π := π) s) at 1. + rewrite lens_ortho_setr_setl. + rewrite lens_get_set. + reflexivity. + Qed. + + Lemma lens_ortho_getr_setl `{Hρgs: !LensGetSet ρ} `{Hρsg: !LensSetGet ρ} s u: + ρ (set π u s) = ρ s. + Proof. + rewrite <- (lens_set_get (π := ρ) s) at 1. + rewrite <- lens_ortho_setr_setl. + rewrite lens_get_set. + reflexivity. + Qed. +End ORTHOGONAL_LENSES. + +Hint Rewrite + @lens_get_set + @lens_set_get + @lens_set_set + @lens_unfold_modify + @lens_eq_set + @lens_set_same_context + @lens_modify_same_context + using typeclasses eauto : lens. + + +(** * Tactics *) + +(** The [autorewrite] base above can be used for normalization, + however it is unable to handle some of the theorems related to + orthogonal lenses. + + For example, we want to use [lens_ortho_getl_setr] to rewrite + occurences of [α (set β v s)], where [α] and [β] are orthogonal + lenses, into simply [α s]. However, our usual trick for combining + [autorewrite] and typeclass resolution does not work, because the + instance of [LensOps α] cannot be obtained by unification when using + [@lens_ortho_getl_setr], which means the premises contain + existentials and we would need [erewrite]. + + Furthermore, we wish to normalize series of nested [set] so that + their order matches the declared instances of [Orthogonal]. However, + rewriting with [@lens_ortho_setr_setl] will only match the first + pair, and fail if they are already in the right order (because a + corresponding instance of [Orthogonal] won't be found). + + So we use the tactic below to apply those explicitely in addition + to using [autorewrite]. *) + +Ltac lens_norm_ortho := + repeat progress + match goal with + | H: context [set ?β ?v (set ?α ?u ?s)] |- _ => + rewrite (lens_ortho_setr_setl u v s) in H + | |- context [set ?β ?v (set ?α ?u ?s)] => + rewrite (lens_ortho_setr_setl u v s) + | H: context [?α (set ?β ?v ?s)] |- _ => + rewrite (lens_ortho_getl_setr α β s v) in H + | |- context [?α (set ?β ?v ?s)] => + rewrite (lens_ortho_getl_setr α β s v) + | H: context [?β (set ?α ?u ?s)] |- _ => + rewrite (lens_ortho_getr_setl α β s u) in H + | |- context [?β (set ?α ?u ?s)] => + rewrite (lens_ortho_getr_setl α β s u) + end. + +Ltac lens_norm := + repeat progress (simpl in *; + lens_norm_ortho; + autorewrite with lens in *). + +Ltac lens_simpl := + repeat progress (lens_norm; autorewrite with lens_simpl in *). + +Ltac lens_unfold := + repeat progress (lens_simpl; unfold set in *). + + +(** * Instances *) + +(** ** Pairs *) + +Section PAIR. + Global Instance fst_lensops A B: LensOps (@fst A B) := { + set a x := (a, snd x) + }. + + Global Instance fst_lens A B: Lens (@fst A B). + Proof. + split; split; intuition. + Qed. + + Global Instance snd_lensops A B: LensOps (@snd A B) := { + set b x := (fst x, b) + }. + + Global Instance snd_lens A B: Lens (@snd A B). + Proof. + split; split; intuition. + Qed. + + (** Here are some product-specific theorems. *) + + Lemma fst_same_context_eq_snd {A B} (x y: A * B): + same_context (@fst A B) x y <-> snd x = snd y. + Proof. + destruct x as [a1 b1]. + destruct y as [a2 b2]. + split; intro H. + * specialize (H a1). + inversion H. + reflexivity. + * intro a. + compute in *. + congruence. + Qed. + + Lemma snd_same_context_eq_fst {A B} (x y: A * B): + same_context (@snd A B) x y <-> fst x = fst y. + Proof. + destruct x as [a1 b1]. + destruct y as [a2 b2]. + split; intro H. + * specialize (H b1). + inversion H. + reflexivity. + * intro b. + compute in *. + congruence. + Qed. +End PAIR. + +Hint Rewrite + @fst_same_context_eq_snd + @snd_same_context_eq_fst + : lens_simpl. + +(** ** Composing lens *) + +Section COMPOSE. + Context {A B C} (π: A -> B) (ρ: B -> C) `{Hπ: Lens _ _ π} `{Hρ: Lens _ _ ρ}. + + Instance compose_lensops: LensOps (compose ρ π) := { + set c a := modify π (set ρ c) a + }. + + Lemma lens_compose_get_simpl a: + compose ρ π a = ρ (π a). + Proof. + reflexivity. + Qed. + + Lemma lens_compose_set_simpl c a: + set (compose ρ π) c a = modify π (set ρ c) a. + Proof. + reflexivity. + Qed. + + Instance lens_compose_get_set: LensGetSet (compose ρ π). + Proof. + constructor; simpl; intros. + rewrite lens_compose_get_simpl. + rewrite lens_compose_set_simpl. + autorewrite with lens. + reflexivity. + Qed. + + Instance lens_compose_set_get: LensSetGet (compose ρ π). + Proof. + constructor; simpl; intros. + rewrite lens_compose_get_simpl. + rewrite lens_compose_set_simpl. + autorewrite with lens. + reflexivity. + Qed. + + Instance lens_compose_set_set: LensSetSet (compose ρ π). + Proof. + constructor; simpl; intros. + rewrite !lens_compose_set_simpl. + autorewrite with lens. + reflexivity. + Qed. + + Instance lens_compose: Lens (compose ρ π) := {}. +End COMPOSE. + +Hint Rewrite + @lens_compose_get_simpl + @lens_compose_set_simpl + using typeclasses eauto : lens_simpl. + +(** ** Initial and terminal objects *) + +Section INITIAL. + Context {V: Type} (π: Empty_set -> V). + + Global Instance lens_empty_ops: LensOps π := { + set a := Empty_set_rect _ + }. + + Global Instance lens_empty: Lens π. + Proof. + split; split; tauto. + Qed. +End INITIAL. + +Section COPRODUCT. + Context {A B V} {α: A -> V} {β: B -> V}. + Context `{αops: LensOps A V α}. + Context `{βops: LensOps B V β}. + + Definition fcoprod (f: A -> V) (g: B -> V): A + B -> V := + fun s => match s with + | inl a => f a + | inr b => g b + end. + + Notation "[ f , g ]" := (fcoprod f g). + + Instance lens_coprod_ops: LensOps [α, β] := { + set v ab := + match ab with + | inl a => inl (set α v a) + | inr b => inr (set β v b) + end + }. + + Instance lens_coprod_get_set: + LensGetSet α -> + LensGetSet β -> + LensGetSet [α,β]. + Proof. + constructor. + intros v [a|b]; + unfold set; simpl. + * lens_norm; reflexivity. + * lens_norm; reflexivity. + Qed. + + Instance lens_coprod_set_get: + LensSetGet α -> + LensSetGet β -> + LensSetGet [α,β]. + Proof. + constructor. + intros [a|b]; + unfold set; simpl. + * lens_norm; reflexivity. + * lens_norm; reflexivity. + Qed. + + Instance lens_coprod_set_set: + LensSetSet α -> + LensSetSet β -> + LensSetSet [α,β]. + Proof. + constructor. + intros u v [a|b]; + unfold set; simpl; + f_equal. + * lens_norm; reflexivity. + * lens_norm; reflexivity. + Qed. + + Global Instance lens_coprod: + Lens α -> + Lens β -> + Lens [α,β]. + Proof. + constructor; + typeclasses eauto. + Qed. +End COPRODUCT. + +Section TERMINAL. + Context {S: Type} (π: S -> unit). + + Instance lens_unit_ops: LensOps (S := S) π := { + set u s := s + }. + + Instance lens_unit: Lens π. + Proof. + assert (forall u v: unit, u = v) by (intros [] []; reflexivity). + firstorder. + Qed. +End TERMINAL. + +(** Interestingly, the product only works for orthogonal lenses. *) + +Section PRODUCT. + Context {S A B} {α: S -> A} {β: S -> B}. + Context `{αops: LensOps S A α}. + Context `{βops: LensOps S B β}. + + Definition fprod (f: S -> A) (g: S -> B): S -> A * B := + fun s => (f s, g s). + + Notation "〈 f , g 〉" := (fprod f g). + + Instance lens_prod_ops: LensOps 〈α, β〉 := { + set ab s := set α (fst ab) (set β (snd ab) s) + }. + + Ltac go := + unfold set, fprod; + simpl; + lens_norm; + try reflexivity. + + Instance lens_prod_get_set: + Orthogonal α β -> + LensSetGet β -> + LensGetSet α -> + LensGetSet β -> + LensGetSet 〈α,β〉. + Proof. + constructor. + intros [a b] s. + go. + Qed. + + Instance lens_prod_set_get: + LensSetGet α -> + LensSetGet β -> + LensSetGet 〈α,β〉. + Proof. + constructor. + intros s. + go. + Qed. + + Instance lens_prod_set_set: + Orthogonal α β -> + LensSetSet α -> + LensSetSet β -> + LensSetSet 〈α,β〉. + Proof. + constructor. + intros [a1 b1] [a2 b2] s. + go. + Qed. + + Global Instance lens_prod: + Orthogonal α β -> + Lens α -> + Lens β -> + Lens 〈α,β〉. + Proof. + constructor; + typeclasses eauto. + Qed. +End PRODUCT. diff --git a/src/lib/Monad/Applicative.v b/src/lib/Monad/Applicative.v new file mode 100755 index 0000000..96ef6e0 --- /dev/null +++ b/src/lib/Monad/Applicative.v @@ -0,0 +1,19 @@ +Set Implicit Arguments. +Set Maximal Implicit Insertion. + +Class Applicative T := +{ pure : forall {A}, A -> T A +; ap : forall {A B}, T (A -> B) -> T A -> T B +}. + +Module ApplicativeNotation. + Notation "f <*> x" := (ap f x) (at level 51, right associativity). +End ApplicativeNotation. +Import ApplicativeNotation. + +Section applicative. + Context {T} {AT:Applicative T}. + + Definition liftA {A B} (f:A -> B) (aT:T A) : T B := pure f <*> aT. + Definition liftA2 {A B C} (f:A -> B -> C) (aT:T A) (bT:T B) : T C := liftA f aT <*> bT. +End applicative. \ No newline at end of file diff --git a/src/lib/Monad/ContOptionMonad.v b/src/lib/Monad/ContOptionMonad.v new file mode 100755 index 0000000..eb9c4bc --- /dev/null +++ b/src/lib/Monad/ContOptionMonad.v @@ -0,0 +1,91 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +Require Import DeepSpec.lib.Monad.Monad. +Require Import DeepSpec.lib.Monad.MonadZero. + +Set Implicit Arguments. +Set Strict Implicit. + +Import MonadNotation. +Open Scope monad_scope. + +(*****************************************) + + +(** * The continuation-option monad + Equivalent to [ContT (Maybe a)] in Haskell *) +Definition ContOpt a := forall r, (a -> option r) -> option r. + + +Global Instance Monad_ContOpt : Monad ContOpt := + { ret A a := fun r k => k a + ; bind A B ma f := fun r k => ma r (fun a => f a r k) + }. + +Global Instance Zero_ContOpt : MonadZero ContOpt := + { mzero _A := fun r k => None }. + +(* Todo: replace uses of these functions with the typeclass ones. *) +Definition Csome {A}(a : A) : ContOpt A := fun r k => k a. +Definition Cnone {A} : ContOpt A := fun r k => None. + +Definition Cguard_sum {P Q}(b : {P} + {Q}) : ContOpt unit := + fun r k => if b then k tt else None. + +Definition Cbind {A B}(ma : ContOpt A)(f : A -> ContOpt B) : ContOpt B := + fun r k => ma r (fun a => f a r k). +Definition Clift {A B}(f : A -> B) : ContOpt A -> ContOpt B := + fun m => Cbind m (fun a => Csome (f a)). + +Definition runContOpt {A}(m : ContOpt A) : option A := m A (@Some A). +Definition optContOpt {A}(a' : option A) : ContOpt A := + fun r k => + match a' with + | Some a => k a + | None => None + end. + +(** - [Csome] is also known as [return] or [unit] + - [Cnone] is also known as [mzero] + - [Cbind] is also known as [bind] + - [optContOpt] is the left and right inverse of [runContOpt]. +*) + +Lemma run_opt_ContOpt_ident {A}(a' : option A) + : runContOpt (optContOpt a') = a'. +Proof. + destruct a'; reflexivity. +Qed. + +(** How to represent that a [ContOpt] computation ruturns a certain value? + [ContOpt_Some_eq m a] is a proposition that [m] will never fail to produce + a result and the return value is [a]. This definition is not used, but the + body can be seen inlined in various places. + + The other direction, with the equality on functions expressed extensionally, + [opt_run_ContOpt_ident_ext : forall A (m : ContOpt A) r k, + optContOpt (runContOpt m) r k = m r k], requires parametricity to prove. *) +Definition ContOpt_Some_eq {A}(m : ContOpt A)(a : A) : Prop := + forall r k, m r k = k a. diff --git a/src/lib/Monad/Functor.v b/src/lib/Monad/Functor.v new file mode 100755 index 0000000..6dc4191 --- /dev/null +++ b/src/lib/Monad/Functor.v @@ -0,0 +1,16 @@ +Set Implicit Arguments. +Set Strict Implicit. + +Section functor. + Variable F : Type -> Type. + + Class Functor : Type := + { fmap : forall A B, (A -> B) -> F A -> F B }. + + Definition ID {T : Type} (f : T -> T) : Prop := + forall x, f x = x. +End functor. + +Module FunctorNotation. + Notation "f <$> x" := (@fmap _ _ f x) (at level 51, right associativity). +End FunctorNotation. diff --git a/src/lib/Monad/Monad.v b/src/lib/Monad/Monad.v new file mode 100755 index 0000000..59ae7ef --- /dev/null +++ b/src/lib/Monad/Monad.v @@ -0,0 +1,61 @@ +Require Import DeepSpec.lib.Monad.Functor. +Require Import DeepSpec.lib.Monad.Applicative. + +Set Implicit Arguments. +Set Strict Implicit. + +Class Monad (m : Type -> Type) : Type := +{ ret : forall {t}, t -> m t +; bind : forall {t u}, m t -> (t -> m u) -> m u +}. + +Section monadic. + Variable m : Type -> Type. + Context {M : Monad m}. + + Definition liftM T U (f : T -> U) : m T -> m U := + fun x => bind x (fun x => ret (f x)). + + Definition liftM2 T U V (f : T -> U -> V) : m T -> m U -> m V := + Eval cbv beta iota zeta delta [ liftM ] in + fun x y => bind x (fun x => liftM (f x) y). + + Definition liftM3 T U V W (f : T -> U -> V -> W) : m T -> m U -> m V -> m W := + Eval cbv beta iota zeta delta [ liftM2 ] in + fun x y z => bind x (fun x => liftM2 (f x) y z). + + Definition apM {A B} (fM:m (A -> B)) (aM:m A) : m B := + bind fM (fun f => liftM f aM). + + Definition bind2 {A B C : Type} (ma : m (A * B)) (f : A -> B -> m C) : m C := + bind ma (fun a => match a with | (x, y) => f x y end). + +End monadic. + +Arguments bind2 {m M A B C} ma f. + +Module MonadNotation. + + Delimit Scope monad_scope with monad. + + Notation "c >>= f" := (@bind _ _ _ _ c f) (at level 50, left associativity) : monad_scope. + Notation "f =<< c" := (@bind _ _ _ _ c f) (at level 51, right associativity) : monad_scope. + + Notation "x <- c1 ;; c2" := (@bind _ _ _ _ c1 (fun x => c2)) + (at level 100, c1 at next level, right associativity) : monad_scope. + + Notation "e1 ;; e2" := (_ <- e1%monad ;; e2%monad)%monad + (at level 100, right associativity) : monad_scope. + + Notation "'unpack' A , B <- y ;; z" := (bind2 y (fun A => (fun B => z))) + (at level 100, A ident, B ident, y at next level, only parsing, right associativity) : monad_scope. + +End MonadNotation. + +Instance Functor_Monad {m} {M:Monad m} : Functor m := + { fmap := @liftM _ _ }. + +Instance Applicative_Monad {m} {M:Monad m} : Applicative m := +{ pure := @ret _ _ +; ap := @apM _ _ +}. diff --git a/src/lib/Monad/MonadExc.v b/src/lib/Monad/MonadExc.v new file mode 100755 index 0000000..56e4bc7 --- /dev/null +++ b/src/lib/Monad/MonadExc.v @@ -0,0 +1,9 @@ +Require Import Monad. + +Set Implicit Arguments. +Set Maximal Implicit Arguments. + +Class MonadExc E (m : Type -> Type) : Type := +{ raise : forall {T}, E -> m T +; catch : forall {T}, m T -> (E -> m T) -> m T +}. diff --git a/src/lib/Monad/MonadInv.v b/src/lib/Monad/MonadInv.v new file mode 100755 index 0000000..6a5a71a --- /dev/null +++ b/src/lib/Monad/MonadInv.v @@ -0,0 +1,101 @@ +Require Import DeepSpec.lib.Monad.Monad. +Require Import DeepSpec.lib.Monad.MonadLaws. + +Set Implicit Arguments. +Set Strict Implicit. + +Section MonadInvDef. + Variable m : Type -> Type. + Variable M : Monad m. + Class MonadInvRet := + monad_inv_ret: + forall {A} (x y: A), + ret x = ret y -> x = y. + + Class MonadInvBind := + monad_inv_bind_extract: + forall {A B} (f: A -> m B) (ma: m A) (b: B), + bind ma f = ret b -> {a | ma = ret a}. + + Class MonadInvBindWeak := + monad_inv_bind_weak: + forall {A B} (f: A -> m B) (ma: m A) (b: B), + bind ma f = ret b -> exists a, f a = ret b. +End MonadInvDef. + +Section MonadInv. + Context `{HM: Monad}. + Context `{HMLaws: !MonadLaws HM}. + Context `{HMbind: !MonadInvBind HM}. + Context `{HMret: !MonadInvRet HM}. + + Global Instance mond_inv_bind_inv_bind_weak: + MonadInvBindWeak HM. + Proof. + intros A B f ma b H. + destruct (monad_inv_bind_extract f ma b H) as [a Ha]. + exists a. subst. + rewrite bind_of_return in H; tauto. + Qed. + + Lemma monad_inv_bind {A B} (f: A -> m B) (ma: m A) (b: B): + bind ma f = ret b -> + {a | ma = ret a /\ f a = ret b}. + Proof. + intros H. + destruct (monad_inv_bind_extract f ma b H) as [a Ha]. + exists a. + split. + * assumption. + * subst. + rewrite bind_of_return in H; tauto. + Qed. + + Lemma monad_inv_bind_iff {A B} (f: A -> m B) (ma: m A) (b: B): + bind ma f = ret b <-> + exists a, ma = ret a /\ f a = ret b. + Proof. + split. + * intros H. + apply monad_inv_bind in H. + destruct H; eauto. + * intros [a [Hma Hfa]]. + subst. + rewrite bind_of_return; tauto. + Qed. + + Lemma monad_inv_bind_inv {A B} {f: A -> m B} {ma: m A} {a: A} {b: B}: + ma = ret a -> + f a = ret b -> + bind ma f = ret b. + Proof. + intros. + apply monad_inv_bind_iff. + exists a; tauto. + Qed. + + Lemma monad_inv_ret_iff {A} (x y: A): + ret x = ret y <-> x = y. + Proof. + split. + apply monad_inv_ret; auto. + apply f_equal. + Qed. + + (* this version is used in the backend part. *) + Lemma bind_some: forall A B (a: m A) (f: A -> m B) (b: B), + bind a f = ret b -> + exists a0, a = ret a0 /\ f a0 = ret b. + Proof. + intros. + destruct (monad_inv_bind_extract _ _ _ H). + eexists. + split. + - eassumption. + - rewrite e in H. + rewrite (@bind_of_return _ _ HMLaws) in H. + assumption. + Qed. + +End MonadInv. + diff --git a/src/lib/Monad/MonadLaws.v b/src/lib/Monad/MonadLaws.v new file mode 100755 index 0000000..3050540 --- /dev/null +++ b/src/lib/Monad/MonadLaws.v @@ -0,0 +1,131 @@ +Require Import FunctionalExtensionality. + +Require Import DeepSpec.lib.Monad.Monad. +Require Import DeepSpec.lib.Monad.MonadTrans. +Require Import DeepSpec.lib.Monad.MonadState. +Require Import DeepSpec.lib.Monad.MonadReader. +Require Import DeepSpec.lib.Monad.MonadZero. + +Set Implicit Arguments. +Set Strict Implicit. + +Section MonadLaws. + Variable m : Type -> Type. + Variable M : Monad m. + + Class MonadLaws : Type := + { bind_of_return : forall A B, + forall (a:A) (f:A -> m B), + (bind (ret a) f) = (f a) + ; return_of_bind : forall A, + forall (aM:m A) (f:A -> m A), + (forall x, (f x) = (ret x)) -> + (bind aM f) = aM + ; bind_associativity : + forall A B C, + forall (aM:m A) (f:A -> m B) (g:B -> m C), + (bind (bind aM f) g) = (bind aM (fun a => bind (f a) g)) + }. + + (* + Add Parametric Morphism T U (ML : MonadLaws) : (@bind _ _ T U) + with signature (eq ==> (eq ==> eq) ==> eq) + as bind_morph. + Proof. intros. + red in H. + firstorder. + eapply bind_proper; auto. Qed. + + Add Parametric Morphism T U (tT : type T) (tU : type U) (tokT : typeOk tT) (tokU : typeOk tU) (ML : MonadLaws) (c : m T) (Pc : proper c) : (@bind _ _ T U c) + with signature ((equal ==> equal) ==> equal) + as bind_1_morph. + Proof. + eapply bind_proper; auto. eapply preflexive; [ | eapply Pc ]. + eapply equiv_prefl; auto. + Qed. + + Add Parametric Morphism T (tT : type T) (tokT : typeOk tT) (ML : MonadLaws) : (@ret _ _ T) + with signature (equal ==> equal) + as ret_morph. + Proof. eapply ret_proper; auto. Qed. + + *) + + Class MonadTLaws (n : Type -> Type) (nM : Monad n) (MT : MonadT m n) : Type := + { lift_ret : forall T, + forall x : T, + (lift (ret x)) = (ret x) + ; lift_bind : forall T U, + forall (c : n T) (f : T -> n U), + (lift (bind c f)) = (bind (lift c) (fun x => lift (f x))) + }. + + + Section with_state. + Context {S : Type}. + + Class MonadStateLaws (MS : MonadState S m) : Type := + { get_put : (bind get put) = (ret tt) + ; put_get : forall x, + (bind (put x) (fun _ => get)) = (bind (put x) (fun _ => ret x)) + ; put_put : forall A, + forall (x y : S) (f : unit -> m A), + (bind (put x) (fun _ => bind (put y) f)) = (bind (put y) f) + ; get_get : forall A, + forall (f : S -> S -> m A), + (bind get (fun s => bind get (f s))) = (bind get (fun s => f s s)) + ; get_ignore : forall A, + forall (aM : m A), + (bind get (fun _ => aM)) = aM + }. + + Class MonadReaderLaws (MR : MonadReader S m) : Type := + { ask_local : forall f, + (local f ask) = (bind ask (fun x => ret (f x))) + ; local_bind : forall A B, + forall aM f (g : A -> m B), + (local f (bind aM g)) = (bind (local f aM) (fun x => local f (g x))) + ; local_ret : forall A, + forall (x : A) f, + (local f (ret x)) = (ret x) + ; local_local : forall T, + forall (s s' : S -> S) (c : m T), + (local s (local s' c)) = (local (fun x => s' (s x)) c) + }. + + End with_state. + + Class MonadZeroLaws (MZ : MonadZero m) : Type := + { bind_zero : forall A B, + forall (f : A -> m B), + (bind mzero f) = mzero + }. + +End MonadLaws. + + + +Hint Rewrite bind_of_return using assumption : monads. +Hint Rewrite bind_associativity using assumption : monads. +Hint Rewrite lift_ret using assumption : monads. +Hint Rewrite lift_bind using assumption : monads. +Hint Rewrite get_put using assumption : monads. +Hint Rewrite put_get using assumption : monads. +Hint Rewrite get_get using assumption : monads. +Hint Rewrite bind_zero using assumption : monads. + +Ltac monad_simpl := +repeat (autorewrite with monads; +try match goal with + | |- context[@bind ?m ?mLaws ?A ?B ?a ?b] => + match b with + | (fun x => _) => + let E := fresh "E" in + let Heq := fresh "H" in + let x := fresh "x" in + evar(E : A -> m B); + assert (b = E) by (unfold E; extensionality x; autorewrite with monads; reflexivity); + unfold E in *; + rewrite Heq; clear Heq + end +end). diff --git a/src/lib/Monad/MonadPlus.v b/src/lib/Monad/MonadPlus.v new file mode 100755 index 0000000..2319023 --- /dev/null +++ b/src/lib/Monad/MonadPlus.v @@ -0,0 +1,17 @@ +Require Import DeepSpec.lib.Monad.Monad. + +Set Implicit Arguments. +Set Maximal Implicit Arguments. + +Class MonadPlus (m : Type -> Type) : Type := +{ mplus : forall {A B:Type}, m A -> m B -> m (A + B)%type }. + +Definition mjoin {m : Type -> Type} {M : Monad m} {MP : MonadPlus m} {T} (a b : m T) : m T := + bind (mplus a b) (fun x => + match x with + | inl x | inr x => ret x + end). + +Module MonadPlusNotation. + Notation "x <+> y" := (@mplus _ _ _ _ x y) (at level 49, right associativity) : monad_scope. +End MonadPlusNotation. diff --git a/src/lib/Monad/MonadReader.v b/src/lib/Monad/MonadReader.v new file mode 100755 index 0000000..b194485 --- /dev/null +++ b/src/lib/Monad/MonadReader.v @@ -0,0 +1,33 @@ +Require Import DeepSpec.lib.Monad.Monad. + +Set Implicit Arguments. +Set Maximal Implicit Arguments. + +Class MonadReader (T : Type) (m : Type -> Type) : Type := +{ local : forall {t}, (T -> T) -> m t -> m t +; ask : m T +}. + +Section monadic. + Variable m : Type -> Type. + Context {M : Monad m}. + Variable T : Type. + Context {MR : MonadReader T m}. + + Definition asks {U} (f : T -> U) : m U := + bind ask (fun x => ret (f x)). + +End monadic. + +Section SubReader. + Variable m : Type -> Type. + Context {M : Monad m}. + Variable T S : Type. + Context {MR : MonadReader T m}. + + Definition ReaderProd (f : T -> S) (g : S -> T -> T) + : MonadReader S m := + {| ask := @asks m M T MR S f + ; local := fun _T up (c : m _T) => @local T m MR _ (fun s => g (up (f s)) s) c + |}. +End SubReader. diff --git a/src/lib/Monad/MonadState.v b/src/lib/Monad/MonadState.v new file mode 100755 index 0000000..e160d14 --- /dev/null +++ b/src/lib/Monad/MonadState.v @@ -0,0 +1,35 @@ +Require Import DeepSpec.lib.Monad.Monad. + +Set Implicit Arguments. +Set Maximal Implicit Arguments. + +Class MonadState (T : Type) (m : Type -> Type) : Type := +{ get : m T +; put : T -> m unit +}. + +Section monadic. + Variable m : Type -> Type. + Context {M : Monad m}. + Variable T : Type. + Context {MS : MonadState T m}. + + Definition modify (f : T -> T) : m unit := + bind get (fun x => put (f x)). + Definition gets {U} (f : T -> U) : m U := + bind get (fun x => ret (f x)). + +End monadic. + +Section SubState. + Variable m : Type -> Type. + Context {M : Monad m}. + Variable T S : Type. + Context {MS : MonadState T m}. + + Definition StateProd (f : T -> S) (g : S -> T -> T) + : MonadState S m := + {| get := @gets m M T MS S f + ; put := fun x => bind get (fun s => put (g x s)) + |}. +End SubState. diff --git a/src/lib/Monad/MonadTrans.v b/src/lib/Monad/MonadTrans.v new file mode 100755 index 0000000..f23b160 --- /dev/null +++ b/src/lib/Monad/MonadTrans.v @@ -0,0 +1,8 @@ +Require Import DeepSpec.lib.Monad.Monad. + +Set Implicit Arguments. +Set Maximal Implicit Arguments. + +Class MonadT (m : Type -> Type) (mt : Type -> Type) : Type := +{ lift : forall {t}, mt t -> m t }. + diff --git a/src/lib/Monad/MonadZero.v b/src/lib/Monad/MonadZero.v new file mode 100755 index 0000000..e33d0b9 --- /dev/null +++ b/src/lib/Monad/MonadZero.v @@ -0,0 +1,17 @@ +Require Import DeepSpec.lib.Monad.Monad. + +Set Implicit Arguments. +Set Maximal Implicit Arguments. + +Class MonadZero (m : Type -> Type) : Type := +{ mzero : forall {T}, m T }. + +Section ZeroFuncs. + Context {m : Type -> Type}. + Context {Monad_m : Monad m}. + Context {Zero_m : MonadZero m}. + + Definition guard (b : bool) : m unit := + if b then ret tt else mzero. + +End ZeroFuncs. \ No newline at end of file diff --git a/src/lib/Monad/OptErrMonad.v b/src/lib/Monad/OptErrMonad.v new file mode 100644 index 0000000..c492786 --- /dev/null +++ b/src/lib/Monad/OptErrMonad.v @@ -0,0 +1,51 @@ +(* This is basically the same thing as in the file EitherMonad.v, + but I wanted to use more explicit constructor names to keep things simple. *) + +Require Import DeepSpec.lib.Monad.Monad. +Require Import DeepSpec.lib.Monad.MonadExc. +Require Import DeepSpec.lib.Monad.MonadInv. +Require Import Coq.Strings.String. + + +Inductive optErr (A:Type) := +| Success : A -> optErr A +| Error : string -> optErr A. + +Arguments Success {A} a. +Arguments Error {A} msg. + +Set Implicit Arguments. +Set Strict Implicit. + +Import MonadNotation. +Open Scope monad_scope. + +Global Instance Monad_optErr : Monad optErr := + { ret := fun _ v => Success v + ; bind := fun _ _ c1 c2 => match c1 with + | Error msg => Error msg + | Success v => c2 v + end + }. + +Global Instance Exception_optErr : MonadExc string optErr := + { raise := fun _ v => Error v + ; catch := fun _ c h => match c with + | Error v => h v + | x => x + end + }. + +Global Instance MonadInvBind_optErr: MonadInvBind Monad_optErr. +Proof. + unfold MonadInvBind. intros. + destruct ma. + - exists a. reflexivity. + - discriminate H. +Defined. + +Global Instance MonadInvRet_optErr: MonadInvRet Monad_optErr. +Proof. + unfold MonadInvRet. simpl. intros. + injection H; auto. +Defined. diff --git a/src/lib/Monad/OptErrMonadLaws.v b/src/lib/Monad/OptErrMonadLaws.v new file mode 100644 index 0000000..5faeca8 --- /dev/null +++ b/src/lib/Monad/OptErrMonadLaws.v @@ -0,0 +1,52 @@ +Require Import FunctionalExtensionality. + +Require Import DeepSpec.lib.Monad.Monad. +Require Import DeepSpec.lib.Monad.MonadTrans. +Require Import DeepSpec.lib.Monad.MonadLaws. +Require Import DeepSpec.lib.Monad.MonadState. +Require Import DeepSpec.lib.Monad.MonadZero. +Require Import DeepSpec.lib.Monad.MonadPlus. + +Require Import DeepSpec.lib.Monad.OptErrMonad. + +Set Implicit Arguments. +Set Strict Implicit. + + +Section Laws. + Variable m : Type -> Type. + Variable Monad_m : Monad m. + Variable ML_m : MonadLaws Monad_m. + + Theorem equal_match : forall (A B : Type), + forall (x y : optErr A) (a b : B) (f g : A -> B), + x = y -> + a = b -> + f = g -> + match x with + | Success a => f a + | Error _ => a + end + = + match y with + | Success a => g a + | Error _ => b + end. + Proof. + destruct x; destruct y; simpl; intros; congruence. + Qed. + + Global Instance MonadLaws_optErr : MonadLaws Monad_optErr. + Proof. + constructor. + { (* bind_of_return *) + intros. simpl. reflexivity. } + { (* return_of_bind *) + simpl; intros. + destruct aM; congruence. } + { (* bind_associativity *) + simpl; intros. + destruct aM; [destruct (f a)|]; congruence. } + Qed. + +End Laws. diff --git a/src/lib/Monad/OptionMonad.v b/src/lib/Monad/OptionMonad.v new file mode 100755 index 0000000..91ddd8b --- /dev/null +++ b/src/lib/Monad/OptionMonad.v @@ -0,0 +1,43 @@ +Require Import DeepSpec.lib.Monad.Monad. +Require Import DeepSpec.lib.Monad.MonadZero. +Require Import DeepSpec.lib.Monad.MonadPlus. +Require Import DeepSpec.lib.Monad.MonadInv. + +Set Implicit Arguments. +Set Strict Implicit. + +Import MonadNotation. +Open Scope monad_scope. + +Global Instance Monad_option : Monad option := +{ ret := @Some +; bind := fun _ _ c1 c2 => match c1 with + | None => None + | Some v => c2 v + end +}. + +Global Instance Zero_option : MonadZero option := +{ mzero := @None }. + +Global Instance Plus_option : MonadPlus option := +{ mplus _A _B aM bM := + match aM with + | None => liftM inr bM + | Some a => Some (inl a) + end +}. + +Global Instance MonadInvBind_option: MonadInvBind Monad_option. +Proof. + unfold MonadInvBind. intros. + destruct ma. + - exists a. reflexivity. + - discriminate H. +Defined. + +Global Instance MonadInvRet_option: MonadInvRet Monad_option. +Proof. + unfold MonadInvRet. simpl. intros. + injection H; auto. +Defined. diff --git a/src/lib/Monad/OptionMonadLaws.v b/src/lib/Monad/OptionMonadLaws.v new file mode 100755 index 0000000..b314f42 --- /dev/null +++ b/src/lib/Monad/OptionMonadLaws.v @@ -0,0 +1,59 @@ +Require Import FunctionalExtensionality. + +Require Import DeepSpec.lib.Monad.Monad. +Require Import DeepSpec.lib.Monad.MonadTrans. +Require Import DeepSpec.lib.Monad.MonadLaws. +Require Import DeepSpec.lib.Monad.MonadState. +Require Import DeepSpec.lib.Monad.MonadZero. +Require Import DeepSpec.lib.Monad.MonadPlus. + +Require Import DeepSpec.lib.Monad.OptionMonad. + +Set Implicit Arguments. +Set Strict Implicit. + + +Section Laws. + Variable m : Type -> Type. + Variable Monad_m : Monad m. + Variable ML_m : MonadLaws Monad_m. + + Theorem equal_match : forall (A B : Type), + forall (x y : option A) (a b : B) (f g : A -> B), + x = y -> + a = b -> + f = g -> + match x with + | Some a => f a + | None => a + end + = + match y with + | Some a => g a + | None => b + end. + Proof. + destruct x; destruct y; simpl; intros; congruence. + Qed. + + Global Instance MonadLaws_option : MonadLaws Monad_option. + Proof. + constructor. + { (* bind_of_return *) + intros. simpl. reflexivity. } + { (* return_of_bind *) + simpl; intros. + destruct aM; congruence. } + { (* bind_associativity *) + simpl; intros. + destruct aM; [destruct (f a)|]; congruence. } + Qed. + + Global Instance MonadZeroLaws_option : MonadZeroLaws Monad_option _. + Proof. + constructor. + { (* bind_zero *) + simpl; intros; reflexivity. } + Qed. + +End Laws. diff --git a/src/lib/Monad/RunStateTInv.v b/src/lib/Monad/RunStateTInv.v new file mode 100755 index 0000000..fed2889 --- /dev/null +++ b/src/lib/Monad/RunStateTInv.v @@ -0,0 +1,334 @@ +Require Import DeepSpec.lib.Monad.StateMonadOption. +Import DeepSpec.lib.Monad.Monad.MonadNotation. + +Require Import MonadInv. +Require Import StateMonadOption. + +Section StateMonadOption. + Context {D : Type}. + Instance MosT : Monad (osT D) := Monad_stateT D Monad_option. + + Lemma gets_osT2 {A} {f: D -> A} {d s: D} {v: A}: + runStateT (gets f) d = ret (v, s) -> + f d = v /\ d = s. + Proof. + intros H. + simpl in H. + injection H. auto. + Qed. + + Lemma get_osT2 {d s v: D}: + runStateT (get) d = (ret (v, s) : option (D*D)) -> + d = v /\ d = s. + Proof. + intros H. + simpl in H. + injection H. auto. + Qed. + + Lemma modify_osT {f: D -> D} {d s: D} {v: unit}: + runStateT (modify f) d = ret (v, s) -> + v = tt /\ f d = s. + Proof. + intros H. + simpl in H. + injection H. auto. + Qed. + + Lemma put_osT {f d s: D} {v: unit}: + runStateT (put f) d = (ret (v, s) : option (unit * D)) -> + v = tt /\ f = s. + Proof. + intros H. + compute in H. + simpl in H. + injection H. auto. + Qed. + + Lemma ret_osT {A} {d s: D} {v u: A}: + runStateT (ret v) d = ret (u, s) -> + v = u /\ d = s. + Proof. + intros H. + simpl in H. + injection H. + auto. + Qed. + + Lemma guard_pure_osT2 {b: bool} {d s: D} {v:unit}: + runStateT (guard b) d = ret (v, s) -> + b = true /\ d = s. + Proof. + intros H. + unfold guard in H. simpl in H. + destruct b; simpl in H. + - injection H; auto. + - discriminate H. + Qed. + + Import MonadNotation. + + +Lemma bind_runStateT : forall {A B} (m1 : osT D A) (m2 : A -> osT D B) d d' v, + runStateT m1 d = Some (v, d') -> + (runStateT (m1 >>= m2) d = runStateT (ret v >>= m2) d') %monad. +Proof. + simpl. + intros. + rewrite H. + reflexivity. +Qed. + +Lemma bind_of_guard_true : forall A (m : unit -> osT D A), + (MonadZero.guard true >>= m)%monad = (m tt). +Proof. +Proof. + simpl. + intros. + destruct (m tt) as [f]. + reflexivity. +Qed. + +Lemma bind_gets_runStateT : forall A B (f : D -> A) (m : A -> osT D B) d, + runStateT (gets f >>= m)%monad d = runStateT (m (f d)) d. +Proof. + reflexivity. +Qed. + +Lemma bind_modify_runStateT : forall B (f : D -> D) (m : unit -> osT D B) d, + runStateT (modify f >>= m)%monad d = runStateT (m tt) (f d). +Proof. + reflexivity. +Qed. + +End StateMonadOption. + + (* For some reason, doing many "destructs" in ltac very slow. (I + guess maybe the type checking algorithm for "match" is trying do + something sophisticated? The runtime seems at least quadratic in + the size of the term.) So instead we provide each lemma as a + function you can "apply", this turns out to be much faster. *) + Lemma split_runStateT_curried : + forall {P : Prop} {D A B : Type} {c1 : osT D A} {c2 : A -> osT D B} {d s : D} {b : B}, + runStateT (x <- c1;; c2 x) d = Some (b, s) -> + (forall (a : A) (m : D), + runStateT c1 d = ret (a, m) -> runStateT (c2 a) m = ret (b, s) -> P) -> + P. + Proof. + intros. + destruct (split_runStateT H) as [? [? [? ?]]]. + eauto. + Qed. + + Lemma get_osT2_curried : + forall {P : Prop} {D : Type} {d s v : D}, + runStateT (get (m:=osT D)) d = ret (v, s) -> + (d = v -> d = s -> P) -> + P. + Proof. + intros. + destruct (get_osT2 H). + eauto. + Qed. + + + Lemma put_osT_curried : forall {P:Prop} {D : Type} {f d s : D} {v : unit}, + runStateT (put (m:=osT D) f) d = ret (v, s) -> + (v = tt -> f = s -> P) -> + P. + Proof. + intros. + destruct (put_osT H). + auto. + Qed. + + Lemma guard_pure_osT2_curried : forall {P:Prop} {D : Type} {b : bool} {d s : D} {v : unit}, + runStateT (guard (m:=osT D) b) d = ret (v, s) -> + (b = true -> d = s -> P) -> + P. + Proof. + intros. + destruct (guard_pure_osT2 H). + auto. + Qed. + + Lemma gets_osT2_curried : + forall {P : Prop} {D A : Type} {f : D -> A} {d s : D} {v : A}, + runStateT (gets (m:=osT D) f) d = ret (v, s) -> + (f d = v -> d = s -> P) -> P. + intros. + destruct (gets_osT2 H). + auto. + Qed. + + Lemma modify_osT_curried : forall {P:Prop} {D : Type} {f : D -> D} {d s : D} {v : unit}, + runStateT (modify (m:=osT D) f) d = ret (v, s) -> + (v = tt -> f d = s -> P) -> + P. + Proof. + intros. + destruct (modify_osT H). + auto. + Qed. + + Lemma ret_osT_curried : forall {P:Prop} {D A : Type} {d s : D} {v u : A}, + runStateT (ret (m:=osT D) v) d = Some (u, s) + -> (v = u -> d = s -> P) -> P. + Proof. + intros. + destruct (ret_osT H). + auto. + Qed. + + + (* The inv_runStateT1 version deals with a single hypothesis from the context, + without re-scanning all of them. *) + Ltac inv_runStateT1 H := + lazymatch type of H with + | (runStateT (bind _ _) _ = ret (_, _)) => + let H1 := fresh H in + let H2 := fresh H in + apply (split_runStateT_curried H); intros ? ? H1 H2; clear H; + inv_runStateT1 H1; + inv_runStateT1 H2 + + | (runStateT (gets _) _ = ret (_, _)) => + apply (gets_osT2_curried H); intros; clear H + | (runStateT get _ = ret (_, _)) => + apply (get_osT2_curried H); intros; clear H + | (runStateT (ret _) _ = ret (_, _)) => + apply (ret_osT_curried H); intros; clear H + | (runStateT (modify _) _ = ret (_, _)) => + apply (modify_osT_curried H); intros; clear H + | (runStateT (put _) _ = ret (_, _)) => + apply (put_osT_curried H); intros; clear H + | (runStateT (guard _) _ = ret (_, _)) => + apply (guard_pure_osT2_curried H); intros; clear H + + (* And the same for Some instead of ret. They are definitionally + equal, but match differently, so we include patterns for both.*) + | (runStateT (bind _ _) _ = Some (_, _)) => + let H1 := fresh H in + let H2 := fresh H in + apply (split_runStateT_curried H); intros ? ? H1 H2; clear H; + inv_runStateT1 H1; + inv_runStateT1 H2 + + | (runStateT (gets _) _ = Some (_, _)) => + apply (gets_osT2_curried H); intros; clear H + | (runStateT get _ = Some (_, _)) => + apply (get_osT2_curried H); intros; clear H + | (runStateT (ret _) _ = Some (_, _)) => + apply (ret_osT_curried H); intros; clear H + | (runStateT (modify _) _ = Some (_, _)) => + apply (modify_osT_curried H); intros; clear H + | (runStateT (put _) _ = Some (_, _)) => + apply (put_osT_curried H); intros; clear H + | (runStateT (guard _) _ = Some (_, _)) => + apply (guard_pure_osT2_curried H); intros; clear H + | _ => idtac + end. + + (* The _branching variant also destruct if-statements. *) + Ltac inv_runStateT1_branching H := + lazymatch type of H with + | (runStateT (bind _ _) _ = ret (_, _)) => + let H1 := fresh H in + let H2 := fresh H in + apply (split_runStateT_curried H); intros ? ? H1 H2; clear H; + inv_runStateT1_branching H1; + inv_runStateT1_branching H2 + | (runStateT (gets _) _ = ret (_, _)) => + apply (gets_osT2_curried H); intros; clear H + | (runStateT get _ = ret (_, _)) => + apply (get_osT2_curried H); intros; clear H + | (runStateT (ret _) _ = ret (_, _)) => + apply (ret_osT_curried H); intros; clear H + | (runStateT (modify _) _ = ret (_, _)) => + apply (modify_osT_curried H); intros; clear H + | (runStateT (put _) _ = ret (_, _)) => + apply (put_osT_curried H); intros; clear H + | (runStateT (guard _) _ = ret (_, _)) => + apply (guard_pure_osT2_curried H); intros; clear H + | (runStateT (if ?b then _ else _) _ = ret _) => + destruct b eqn:?; + inv_runStateT1_branching H + + (* And the same for Some instead of ret. They are definitionally + equal, but match differently, so we include patterns for both.*) + | (runStateT (bind _ _) _ = Some (_, _)) => + let H1 := fresh H in + let H2 := fresh H in + apply (split_runStateT_curried H); intros ? ? H1 H2; clear H; + inv_runStateT1_branching H1; + inv_runStateT1_branching H2 + | (runStateT (gets _) _ = Some (_, _)) => + apply (gets_osT2_curried H); intros; clear H + | (runStateT get _ = Some (_, _)) => + apply (get_osT2_curried H); intros; clear H + | (runStateT (ret _) _ = Some (_, _)) => + apply (ret_osT_curried H); intros; clear H + | (runStateT (modify _) _ = Some (_, _)) => + apply (modify_osT_curried H); intros; clear H + | (runStateT (put _) _ = Some (_, _)) => + apply (put_osT_curried H); intros; clear H + | (runStateT (guard _) _ = Some (_, _)) => + apply (guard_pure_osT2_curried H); intros; clear H + | (runStateT (if ?b then _ else _) _ = ret _) => + destruct b eqn:?; + inv_runStateT1_branching H + | _ => idtac + end. + + +(* strictly speaking the option monad is not really realated to RunStateT, + but it is used in the stmt_vc, so it is helpful to make invRunStateT + invert it also. *) + +Lemma split_bind_option {A B} {c1: option A} {c2: A -> option B} {b: B}: + (@bind option OptionMonad.Monad_option _ _ c1 c2) = ret b -> + exists a, c1 = ret a /\ (c2 a) = ret b. +Proof. + destruct c1. + + simpl. + eexists; split; auto. + + simpl. inversion 1. +Qed. + + Ltac inv_runStateT := + repeat match goal with + | [ H: (@bind option OptionMonad.Monad_option (_ * _) _ _ _) = ret _ |- _] => destruct (split_bind_option H) as [[? ?] [? ?]]; clear H + + | [ H: (@bind option OptionMonad.Monad_option _ _ _ _) = ret _ |- _] => destruct (split_bind_option H) as [? [? ?]]; clear H + + | [ H : runStateT _ _ = ret (_, _) |- _] => inv_runStateT1 H + | [ H : execStateT _ _ = ret _ |- _] => apply execStateT_runStateT in H; destruct H as [? ?] + | [ H : evalStateT _ _ = ret _ |- _] => apply evalStateT_runStateT in H; destruct H as [? ?] + | [ H : evalStateT mzero _ = ret _ |- _] => discriminate H + + (* The Some and ret constructors are definitionally equal, let's just match both.. *) + | [ H : runStateT _ _ = Some (_, _) |- _] => inv_runStateT1 H + | [ H : execStateT _ _ = Some _ |- _] => apply execStateT_runStateT in H; destruct H as [? ?] + | [ H : evalStateT _ _ = Some _ |- _] => apply evalStateT_runStateT in H; destruct H as [? ?] + | [ H : evalStateT mzero _ = Some _ |- _] => discriminate H + + end. + + Ltac inv_runStateT_branching := + repeat match goal with + | [ H: (@bind option OptionMonad.Monad_option (_ * _) _ _ _) = ret _ |- _] => destruct (split_bind_option H) as [[? ?] [? ?]]; clear H + + | [ H: (@bind option OptionMonad.Monad_option _ _ _ _) = ret _ |- _] => destruct (split_bind_option H) as [? [? ?]]; clear H + + | [ H : runStateT _ _ = ret (_, _) |- _] => inv_runStateT1_branching H + | [ H : execStateT _ _ = ret _ |- _] => apply execStateT_runStateT in H; destruct H as [? ?] + | [ H : evalStateT _ _ = ret _ |- _] => apply evalStateT_runStateT in H; destruct H as [? ?] + | [ H : evalStateT mzero _ = ret _ |- _] => discriminate H + + (* The Some and ret constructors are definitionally equal, let's just match both.. *) + | [ H : runStateT _ _ = Some (_, _) |- _] => inv_runStateT1_branching H + | [ H : execStateT _ _ = Some _ |- _] => apply execStateT_runStateT in H; destruct H as [? ?] + | [ H : evalStateT _ _ = Some _ |- _] => apply evalStateT_runStateT in H; destruct H as [? ?] + | [ H : evalStateT mzero _ = Some _ |- _] => discriminate H + + end. diff --git a/src/lib/Monad/StateMonad.v b/src/lib/Monad/StateMonad.v new file mode 100755 index 0000000..5bf3719 --- /dev/null +++ b/src/lib/Monad/StateMonad.v @@ -0,0 +1,123 @@ +Require Import DeepSpec.lib.Monad.Monad. +Require Import DeepSpec.lib.Monad.MonadTrans. +Require Import DeepSpec.lib.Monad.MonadState. +Require Import DeepSpec.lib.Monad.MonadZero. + +Set Implicit Arguments. +Set Strict Implicit. + +Section StateType. + Variable S : Type. + + Record state (t : Type) : Type := mkState + { runState : S -> t * S }. + + Definition evalState {t} (c : state t) (s : S) : t := + fst (runState c s). + + Definition execState {t} (c : state t) (s : S) : S := + snd (runState c s). + + + Global Instance Monad_state : Monad state := + { ret := fun _ v => mkState (fun s => (v, s)) + ; bind := fun _ _ c1 c2 => + mkState (fun s => + let (v,s) := runState c1 s in + runState (c2 v) s) + }. + + Global Instance MonadState_state : MonadState S state := + { get := mkState (fun x => (x,x)) + ; put := fun v => mkState (fun _ => (tt, v)) + }. + + Variable m : Type -> Type. + + Record stateT (t : Type) : Type := mkStateT + { runStateT : S -> m (t * S)%type }. + + Variable M : Monad m. + + Definition evalStateT {t} (c : stateT t) (s : S) : m t := + bind (runStateT c s) (fun x => ret (fst x)). + + Definition execStateT {t} (c : stateT t) (s : S) : m S := + bind (runStateT c s) (fun x => ret (snd x)). + + + Global Instance Monad_stateT : Monad stateT := + { ret := fun _ x => mkStateT (fun s => @ret _ M _ (x,s)) + ; bind := fun _ _ c1 c2 => + mkStateT (fun s => + @bind _ M _ _ (runStateT c1 s) (fun vs => + let (v,s) := vs in + runStateT (c2 v) s)) + }. + + Global Instance MonadState_stateT : MonadState S stateT := + { get := mkStateT (fun x => ret (x,x)) + ; put := fun v => mkStateT (fun _ => ret (tt, v)) + }. + + Global Instance MonadT_stateT : MonadT stateT m := + { lift := fun _ c => mkStateT (fun s => bind c (fun t => ret (t, s))) + }. + + Global Instance State_State_stateT T (MS : MonadState T m) : MonadState T stateT := + { get := lift get + ; put := fun x => lift (put x) + }. + + Global Instance MonadZero_stateT (MR : MonadZero m) : MonadZero stateT := + { mzero _A := lift mzero + }. + + + (* Todo later. *) + (* + Global Instance MonadReader_stateT T (MR : MonadReader T m) : MonadReader T stateT := + { ask := mkStateT (fun s => bind ask (fun t => ret (t, s))) + ; local := fun _ f c => mkStateT (fun s => local f (runStateT c s)) + }. + + Global Instance MonadWriter_stateT T (Mon : Monoid T) (MR : MonadWriter Mon m) : MonadWriter Mon stateT := + { tell := fun x => mkStateT (fun s => bind (tell x) (fun v => ret (v, s))) + ; listen := fun _ c => mkStateT (fun s => bind (listen (runStateT c s)) + (fun x => let '(a,s,t) := x in + ret (a,t,s))) + ; pass := fun _ c => mkStateT (fun s => bind (runStateT c s) (fun x => + let '(a,t,s) := x in ret (a, s))) + }. + + Global Instance Exc_stateT T (MR : MonadExc T m) : MonadExc T stateT := + { raise := fun _ e => lift (raise e) + ; catch := fun _ body hnd => + mkStateT (fun s => catch (runStateT body s) (fun e => runStateT (hnd e) s)) + }. + + + Global Instance MonadFix_stateT (MF : MonadFix m) : MonadFix stateT := + { mfix := fun _ _ r v => + mkStateT (fun s => mfix2 _ (fun r v s => runStateT (mkStateT (r v)) s) v s) + }. + + Global Instance MonadPlus_stateT (MP : MonadPlus m) : MonadPlus stateT := + { mplus _A _B a b := + mkStateT (fun s => bind (mplus (runStateT a s) (runStateT b s)) + (fun res => match res with + | inl (a,s) => ret (inl a, s) + | inr (b,s) => ret (inr b, s) + end)) + }. +*) + +End StateType. + +Arguments evalState {S} {t} (c) (s). +Arguments execState {S} {t} (c) (s). +Arguments evalStateT {S} {m} {M} {t} (c) (s). +Arguments execStateT {S} {m} {M} {t} (c) (s). +(* +Arguments MonadWriter_stateT {S} {m} {_} {T} {Mon} (_). +*) diff --git a/src/lib/Monad/StateMonadLaws.v b/src/lib/Monad/StateMonadLaws.v new file mode 100755 index 0000000..efab70e --- /dev/null +++ b/src/lib/Monad/StateMonadLaws.v @@ -0,0 +1,154 @@ +Require Import FunctionalExtensionality. + +Require Import DeepSpec.lib.Monad.Monad. +Require Import DeepSpec.lib.Monad.MonadTrans. +Require Import DeepSpec.lib.Monad.MonadLaws. +Require Import DeepSpec.lib.Monad.MonadState. +Require Import DeepSpec.lib.Monad.MonadZero. + +Require Import StateMonad. + +Set Implicit Arguments. +Set Strict Implicit. + +Section Laws. + Variable S : Type. + Variable m : Type -> Type. + Variable Monad_m : Monad m. + Variable ML_m : MonadLaws Monad_m. + + + Lemma equal_runStateT : forall t (m1 m2 : stateT S m t) , + (forall s, runStateT m1 s = runStateT m2 s) -> m1 = m2. + Proof. + intros. + destruct m1. destruct m2. + simpl in H. + f_equal. + extensionality s. + auto. + Qed. + + Global Instance MonadLaws_stateT : MonadLaws (@Monad_stateT S _ Monad_m). + Proof. + constructor. + {(* bind_of_return *) + intros. + apply equal_runStateT. + intros. + simpl. + rewrite bind_of_return by auto. + reflexivity. } + { (* return_of_bind *) + simpl; intros. + apply equal_runStateT. + simpl; intros. + rewrite return_of_bind. + - reflexivity. + - auto. + - intros x. + destruct x. + rewrite H. + reflexivity. } + { (* bind_associativity *) + simpl; intros. + apply equal_runStateT. + intros; simpl. + repeat rewrite bind_associativity by auto. + f_equal. + extensionality a. + destruct a. + reflexivity. } + Qed. + + Global Instance MonadTLaws_stateT : MonadTLaws _ _ (@MonadT_stateT S _ Monad_m). + Proof. + constructor. + { simpl. intros. + apply equal_runStateT. + intros; simpl. + rewrite bind_of_return; auto. } + { simpl. intros. + apply equal_runStateT. + intros; simpl. + rewrite bind_associativity by auto. + rewrite bind_associativity by auto. + f_equal. + extensionality a. + rewrite bind_of_return by auto. + reflexivity. } + Qed. + + +Global Instance MoandStateLaws_stateT : MonadStateLaws (@Monad_stateT S _ Monad_m) (@MonadState_stateT S _ Monad_m). + Proof. + constructor. + { (* get_put *) + simpl. apply equal_runStateT; intros; simpl. + rewrite bind_of_return by auto. + reflexivity. } + { (* put_get *) + simpl; intros; apply equal_runStateT; intros; simpl. + repeat rewrite bind_of_return by auto. + reflexivity. } + { (* put_put *) + simpl; intros; apply equal_runStateT; intros; simpl. + repeat rewrite bind_of_return by auto. + reflexivity. } + { (* get_get *) + simpl; intros; apply equal_runStateT; intros; simpl. + repeat rewrite bind_of_return by auto. + reflexivity. } + {(* get_ignore*) + simpl; intros; apply equal_runStateT; intros; simpl. + repeat rewrite bind_of_return by auto. + reflexivity. } + Qed. + + (* + Global Instance MonadReaderLaws_optionT (s : Type) (t : type s) (tT : typeOk t) (Mr : MonadReader s m) (MLr : MonadReaderLaws Monad_m _ _ Mr) : MonadReaderLaws _ _ _ (@Reader_optionT _ _ _ Mr). + Proof. + constructor. + { simpl. unfold optionT_eq; simpl; intros; unfold liftM. + rewrite local_bind; eauto with typeclass_instances. + (erewrite bind_proper; [ | | | | eapply ask_local | ]); eauto with typeclass_instances. + rewrite bind_associativity; eauto with typeclass_instances. + rewrite bind_associativity; eauto with typeclass_instances. + type_tac. 6: eapply preflexive. + repeat rewrite bind_of_return; eauto with typeclass_instances. + rewrite local_ret; eauto with typeclass_instances. type_tac. + type_tac. eapply equal_match; eauto with typeclass_instances; type_tac. + apply proper_fun; intros. repeat rewrite local_ret; eauto with typeclass_instances. + type_tac; eauto with typeclass_instances. type_tac. + type_tac. eapply equal_match; eauto with typeclass_instances; type_tac. + type_tac. + apply proper_fun; intros. repeat rewrite local_ret; eauto with typeclass_instances. + type_tac. eauto with typeclass_instances. + type_tac. type_tac. } + { simpl. unfold optionT_eq; simpl; intros; unfold liftM. + rewrite local_bind; eauto with typeclass_instances. + type_tac. + destruct x; destruct y; try solve [ inversion H4 ]; type_tac. + inversion H4; assumption. + rewrite local_ret; eauto with typeclass_instances; type_tac. + type_tac. eapply equal_match; eauto with typeclass_instances; type_tac. } + { simpl. unfold optionT_eq; simpl; intros; unfold liftM. + rewrite local_ret; eauto with typeclass_instances; type_tac. } + { simpl. unfold optionT_eq; simpl; intros; unfold liftM. + rewrite local_local; eauto with typeclass_instances; type_tac. } + { unfold local; simpl; intros. red. red. intros. red in H0. + red; simpl. type_tac. } + { Opaque lift. unfold ask; simpl; intros. red. type_tac. + eapply lift_proper; eauto with typeclass_instances. Transparent lift. } + Qed. + *) + + Global Instance MonadZeroLaws_optionT (MZ: MonadZero m) (MZL: MonadZeroLaws _ MZ) : MonadZeroLaws (@Monad_stateT S _ Monad_m) _. + Proof. + constructor. + { simpl; intros; apply equal_runStateT; intros; simpl. + repeat rewrite bind_zero by auto. + reflexivity. } + Qed. + +End Laws. diff --git a/src/lib/Monad/StateMonadOption.v b/src/lib/Monad/StateMonadOption.v new file mode 100755 index 0000000..8ae7c24 --- /dev/null +++ b/src/lib/Monad/StateMonadOption.v @@ -0,0 +1,457 @@ +Require Export DeepSpec.lib.Monad.Monad. +Require Export DeepSpec.lib.Monad.MonadLaws. +Require Export DeepSpec.lib.Monad.MonadInv. +Require Export DeepSpec.lib.Monad.StateMonad. +Require Export DeepSpec.lib.Monad.OptionMonad. +Require Export DeepSpec.lib.Monad.OptionMonadLaws. +Require Export DeepSpec.lib.Monad.MonadZero. +Require Export DeepSpec.lib.Monad.MonadState. +Require Export DeepSpec.lib.Monad.StateMonadLaws. + +Set Implicit Arguments. +Set Strict Implicit. + +Import MonadNotation. +Open Scope monad_scope. + + +Ltac dest_opt opt opt_new l := + let p := fresh "p" in + let opt1 := fresh opt in + destruct (monad_inv_bind _ _ _ opt) as (p & opt1 & opt_new); + clear opt; + rename opt1 into opt; + destruct p as l. + +Tactic Notation "dest_opt" hyp(opt) "as" + ident(opt_new) simple_intropattern(l) := + dest_opt opt opt_new l. +(* +Ltac inj_state H := + match type of H with + | runStateT _ _ = _ => injection H as <- <- + | execStateT _ _ = _ => injection H as <- + | evalStateT _ _ = _ => injection H as <- + | ret (fst _) = _ => injection H as <- + | Some (fst _) = _ => injection H as <- + | ret (snd _) = _ => injection H as <- + | Some (snd _) = _ => injection H as <- + | _ = ret (_,_) => injection H as <- <- + | _ = Some (_,_) => injection H as <- <- + end; + clear H. +*) +Section StateMonadOption. + Variable D : Type. + Definition osT := @stateT D option. + Instance MosT : Monad osT := Monad_stateT D Monad_option. + + Instance MonadState_osT : MonadState D osT. + Proof. + apply MonadState_stateT. + apply Monad_option. + Defined. + + Instance MonadZero_osT : MonadZero osT. + Proof. + apply MonadZero_stateT. + apply Monad_option. + apply Zero_option. + Defined. + + Instance MonadLaws_MosT : MonadLaws MosT. + Proof. + apply MonadLaws_stateT. + apply MonadLaws_option. + Defined. + + Lemma ret_inj {A} {v1 v2:A}: + @ret _ Monad_option _ v1 = ret v2 -> + v1 = v2. + Proof. + intro H. injection H as <-. + reflexivity. + Qed. + + Lemma ret_Some {A} {v1 v2:A}: + @ret _ Monad_option _ v1 = Some v2 -> + v1 = v2. + Proof. + apply ret_inj. + Qed. + + Lemma Some_ret {A} {v1 v2:A}: + Some v1 = @ret _ Monad_option _ v2 -> + v1 = v2. + Proof. + apply ret_inj. + Qed. + + Lemma runStateT_execStateT {A} {c: osT A} {d v s}: + runStateT c d = ret (v, s) -> + execStateT c d = ret s. + Proof. + intros. + simpl; rewrite H; reflexivity. + Qed. + + Lemma runStateT_execStateT' {A} {c: osT A} {d s}: + (exists v, runStateT c d = ret (v, s)) -> + execStateT c d = ret s. + Proof. + intros [v H]. + apply (runStateT_execStateT H). + Qed. + + Lemma runStateT_evalStateT {A} {c: osT A} {d v s}: + runStateT c d = ret (v, s) -> + evalStateT c d = ret v. + Proof. + intros. + simpl; rewrite H; reflexivity. + Qed. + + Lemma runStateT_evalStateT' {A} {c: osT A} {d v}: + (exists s, runStateT c d = ret (v, s)) -> + evalStateT c d = ret v. + Proof. + intros [s H]. + apply (runStateT_evalStateT H). + Qed. + + Lemma execStateT_runStateT {A} {c: osT A} {d s}: + execStateT c d = ret s -> + exists v, runStateT c d = ret (v, s). + Proof. + intros opt. + dest_opt opt as opt' [v' s']. + apply ret_inj in opt'. cbv in opt'. subst s'. + exists v'. assumption. + Qed. + + Lemma evalStateT_runStateT {A} {c: osT A} {d v}: + evalStateT c d = ret v -> + exists s, runStateT c d = ret (v, s). + Proof. + intros opt. + dest_opt opt as opt' [v' s']. + apply ret_inj in opt'. cbv in opt'. subst v'. + exists s'. assumption. + Qed. + + (** pure *) + + Lemma gets_pure_osT {A} {f: D -> A} {d s: D} {v: A}: + runStateT (@gets osT _ _ _ _ f) d = ret (v, s) -> + d = s. + Proof. + intros H. + simpl in H. + injection H. auto. + Qed. + + Lemma guard_pure_osT {b: bool} {d s: D} {v:unit}: + runStateT (guard b) d = ret (v, s) -> + d = s. + Proof. + intros H. + unfold guard in H. simpl in H. + destruct b; simpl in H. + - injection H; auto. + - discriminate H. + Qed. + + (** bind *) + + Lemma bind_osT {A B} {ma: osT A} {f: A -> osT B} {a: A} {b: B} {d s m: D}: + runStateT ma d = ret (a, s) -> + runStateT (f a) s = ret (b, m) -> + runStateT (x <- ma;; f x) d = ret (b, m). + Proof. + intros opta optb. + apply (monad_inv_bind_inv opta optb). + Qed. + + Lemma bind_osT' {A B} {ma: osT A} {f: A -> osT B} {a: A} {b: B} {d s m: D}: + runStateT ma d = ret (a, s) /\ + ( runStateT ma d = ret (a, s) -> runStateT (f a) s = ret (b, m)) -> + runStateT (x <- ma;; f x) d = ret (b, m). + Proof. + intros [opta optb]. + specialize (optb opta). + apply (bind_osT opta optb). + Qed. + + Lemma bind_rewrite_osT_mzero {A B} {c: osT A} {d: D} {f: A -> osT B}: + runStateT c d = mzero -> + runStateT (x<-c;;f x) d = mzero. + Proof. + intros opt. + simpl. rewrite opt. + reflexivity. + Qed. + + Lemma bind_to_none_osT {A B} {c: osT A} {d: D} {f: A -> osT B}: + runStateT (c >>= f) d = mzero -> + {runStateT c d = mzero} + + {exists v s, runStateT c d = ret (v, s) /\ runStateT (f v) s = mzero}. + Proof. + intros opt. + simpl in opt. + destruct (runStateT c d); [right|left]. + - destruct p as [v s]. + exists v, s. auto. + - reflexivity. + Qed. + + Lemma bind_rewrite_osT_runStateT {A B} {d1 d2: D} + {c1 c2: osT A} {c3: A -> osT B}: + runStateT c1 d1 = runStateT c2 d2 -> + runStateT (x<-c1;;c3 x) d1 = runStateT (x<-c2;;c3 x) d2. + Proof. + intros opt. + simpl. + rewrite opt. reflexivity. + Qed. + + Lemma bind_rewrite_osT_ret {A B} {d1 d2: D} {v:A} + {c1: osT A} {c3: A -> osT B}: + runStateT c1 d1 = ret (v, d2) -> + runStateT (x<-c1;;c3 x) d1 = runStateT (c3 v) d2. + Proof. + intros opt. + simpl. + rewrite opt. reflexivity. + Qed. + + Lemma split_runStateT {A B} {c1: osT A} {c2: A -> osT B} {d s: D} {b: B}: + runStateT (x<-c1;; c2 x) d = ret (b, s) -> + exists a m, runStateT c1 d = ret (a, m) /\ runStateT (c2 a) m = ret (b, s). + Proof. + intros H. + destruct (monad_inv_bind _ _ _ H) as [[a m] [H1 H2]]. + exists a, m. + auto. + Qed. + + Lemma split_execStateT {A B} {c1: osT A} {c2: osT B} {d s: D}: + execStateT (c1;; c2) d = ret s -> + exists m, execStateT c1 d = ret m /\ execStateT c2 m = ret s. + Proof. + intros H. + destruct (execStateT_runStateT H) as [b H']; clear H; rename H' into H. + destruct (split_runStateT H) as (a & m & H1 & H2). + exists m. + split. + - apply (runStateT_execStateT H1). + - apply (runStateT_execStateT H2). + Qed. + + (** get *) + Lemma get_osT {d v s: D}: + runStateT get d = ret (v, s) <-> + v = d /\ d = s. + Proof. + simpl. + split. + - intros H. injection H as <- <-. tauto. + - intros [? ?]. subst. reflexivity. + Qed. + + (** put *) + Lemma put_osT {d d' s: D} {v: unit}: + runStateT (put d') d = ret (v, s) <-> + v = tt /\ d' = s. + Proof. + simpl. + split. + - intros opt. injection opt as <- <-. tauto. + - intros [H1 H2]. subst. reflexivity. + Qed. + + (** gets *) + Lemma gets_osT {A} {f: D -> A} {d s: D} {v: A}: + runStateT (@gets osT _ _ _ _ f) d = ret (v, s) <-> + f d = v /\ d = s. + Proof. + simpl. + split. + - intros opt. injection opt as <- <-. + tauto. + - intros [H1 H2]. subst. reflexivity. + Qed. + + Lemma gets_osT_eval {A} {f: D -> A} {d: D} {v: A}: + evalStateT (@gets osT _ _ _ _ f) d = ret v <-> + f d = v. + Proof. + split. + - intros opt. + simpl in opt. + injection opt. + intro Hfeq. + cut (Some (f d, d) = Some (v, d)). + + intro H. injection H. auto. + + injection opt as <-. reflexivity. + - intros Hf. + simpl. rewrite Hf. reflexivity. + Qed. + + (** modify *) + Lemma modify_osT {f : D -> D} {d s: D} {v: unit}: + runStateT (@modify osT _ _ _ f) d = ret (v, s) <-> + v = tt /\ f d = s. + Proof. + simpl. + split. + - intros opt. injection opt as <- <-. auto. + - intros [H1 H2]. subst. reflexivity. + Qed. + + (** guard *) + Lemma guard_osT {b: bool} {m v s}: + runStateT (guard b) m = ret (v, s) <-> + b = true /\ v = tt /\ m = s. + Proof. + unfold guard. + destruct b. + - (* b = true *) + simpl. + split. + + intros opt. injection opt as <- <-. tauto. + + intros (? & ? & ?). subst. reflexivity. + - (* b = false *) + simpl. + split. + + discriminate. + + intros (? & ? & ?). discriminate. + Qed. +(* + Lemma gets_guard_osT' {A} {f: D -> A} {g: A -> bool} {d s: D} {v: unit}: + runStateT (x <- gets f;; guard (g x)) d = ret (v, s) -> + g (f d) = true. + Proof. + unfold guard. + simpl. + destruct (g (f d)). + - reflexivity. + - simpl. discriminate. + Qed. + + Lemma gets_guard_pure_osT {A} {f: D -> A} {g: A -> bool} {d s: D} {v: unit}: + runStateT (x <- gets f;; guard (g x)) d = ret (v, s) -> + d = s. + Proof. + intro H. + destruct (split_runStateT H) as (v' & s' & H1 & H2). + apply gets_pure_osT in H1. + apply guard_pure_osT in H2. + subst. reflexivity. + Qed. + + Lemma gets_guard_osT {A} {f: D -> A} {g: A -> bool} {d s: D} {v: unit}: + runStateT (x <- gets f;; guard (g x)) d = ret (v, s) -> + g (f s) = true. + Proof. + intro H. + assert (d = s). + { apply (gets_guard_pure_osT H). } + rewrite <- H0. + apply (gets_guard_osT' H). + Qed. +*) + + Lemma ret_osT {A} {a: A} {m s: D} {v}: + runStateT (ret a) m = ret (v, s) <-> + a = v /\ m = s. + Proof. + simpl. + split. + - intros H. injection H as <- <-. tauto. + - intros [? ?]. subst. reflexivity. + Qed. + + Lemma match_osT {A} {e: option A} {d s : D} {v: A}: + runStateT + match e with + | Some a => ret a + | None => mzero + end d = ret (v, s) <-> + e = Some v /\ d = s. + Proof. + destruct e; simpl. + - (* e = Some a *) + split. + + intros opt. injection opt as <- <-. tauto. + + intros [H1 H2]. injection H1 as <-. subst. reflexivity. + - (* e = None *) + split. + + discriminate. + + intros [? ?]. discriminate. + Qed. +(* + Lemma eval_ret_osT {A} {a: A} {m:D} {v}: + @evalStateT D option Monad_option A (ret a) m = ret v -> + a = v. + Proof. + intro H. + destruct (evalStateT_runStateT H) as [s Hrun]. + apply run_ret_osT in Hrun. + tauto. + Qed. +*) +End StateMonadOption. + +Ltac exec2run_util H l := + let H' := fresh in + destruct (execStateT_runStateT H) as [l H']; + clear H; rename H' into H. + +Tactic Notation "exec2run" hyp(H) "as" simple_intropattern(l) := + exec2run_util H l. + +Tactic Notation "exec2run" := + apply runStateT_execStateT'. + +Ltac eval2run_util H l := + let H' := fresh in + destruct (evalStateT_runStateT H) as [l H']; + clear H; rename H' into H. + +Tactic Notation "eval2run" hyp(H) "as" simple_intropattern(l) := + eval2run_util H l. + +Tactic Notation "eval2run" := + apply runStateT_evalStateT'. + +Tactic Notation "run2exec" hyp(H) := + apply runStateT_execStateT in H. + +Tactic Notation "run2eval" hyp(H) := + apply runStateT_evalStateT in H. + +(** rewrite runStateT c1 d = _ in + runStateT (x<-c1;;c2 x) d +*) +Ltac bind_rewrite_osT H := + match type of H with + | runStateT _ _ = ret (_, _) => erewrite (bind_rewrite_osT_ret H) + | runStateT _ _ = mzero => erewrite (bind_rewrite_osT_mzero H) + | runStateT _ _ = runStateT _ _ => erewrite (bind_rewrite_osT_runStateT H) + end. + +Ltac split_run_util H l := + apply split_runStateT in H; + destruct H as l. + +Tactic Notation "split_run" hyp(H) "as" simple_intropattern(l) := + split_run_util H l. + +Lemma extract_if_rev {b : bool}: + (if b then false else true) = true <-> + b = false. +Proof. + destruct b. + - (* b = true *) split; auto. + - (* b = false *) tauto. +Qed. diff --git a/src/lib/OProp.v b/src/lib/OProp.v new file mode 100755 index 0000000..b2d4414 --- /dev/null +++ b/src/lib/OProp.v @@ -0,0 +1,288 @@ +(* *********************************************************************) +(* DeepSpec, the language of certified softwares *) +(* *) +(* Shu-Chun Weng, Yale University *) +(* *) +(* Copyright (c) 2013-2015 Shu-Chun Weng . *) +(* *) +(* This program is free software; you can redistribute it and/or *) +(* modify it under the terms of the GNU General Public License *) +(* version 2 as published by the Free Software Foundation. Note that *) +(* the only valid version of the GPL for this work is version 2, not *) +(* v2.2 or v3.x or whatever. *) +(* *) +(* This program is distributed in the hope that it will be useful, *) +(* but WITHOUT ANY WARRANTY; without even the implied warranty of *) +(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *) +(* GNU General Public License for more details. *) +(* *) +(* You should have received a copy of the GNU General Public License *) +(* along with this program; if not, write to the Free Software *) +(* Foundation, Inc., 51 Franklin Street, Fifth Floor, Boston, *) +(* MA 02110-1301 USA. *) +(* *********************************************************************) + +(** * Optional proposition + Many constructs do not need extra side conditions, but [True] polutes the + final result regardless. [OProp] replaces [Prop] and, along with [oand], + only inject to Coq [Prop] at the end, without redundent [True]s. *) +Definition OProp := option Prop. +Definition otrue : OProp := None. +Definition ofalse : OProp := Some False. +Definition oprop : Prop -> OProp := @Some Prop. + +Definition oProp (op : OProp) := match op with + | None => True + | Some p => p + end. + +Definition oand (op1 op2 : OProp) := match op1, op2 with + | None, _ => op2 + | _, None => op1 + | Some p1, Some p2 => Some (p1 /\ p2) + end. + +Lemma oand_distr op1 op2 : oProp (oand op1 op2) <-> oProp op1 /\ oProp op2. +Proof. + destruct op1 as [ pp1 |], op2 as [ pp2 |]; simpl; + split; try intros [ p1 p2 ]; auto. +Qed. + +Lemma oconj op1 op2 (p1 : oProp op1)(p2 : oProp op2) : oProp (oand op1 op2). +Proof (proj2 (oand_distr _ _) (conj p1 p2)). + +Definition OProp1 A := option (A -> Prop). +Definition otrue1 {A} : OProp1 A := None. +Definition ofalse1 {A} : OProp1 A := Some (fun _ => False). +Definition oprop1 {A} : (A -> Prop) -> OProp1 A := @Some (A -> Prop). + +Definition oProp1 {A}(op : OProp1 A) := match op with + | None => fun _ => True + | Some p => p + end. + +Definition omap1 {A A'} f (op : OProp1 A) : OProp1 A' := + match op with + | None => None + | Some p => Some (fun a => f p a) + end. + +Definition oand1 {A}(op1 op2 : OProp1 A) := match op1, op2 with + | None, _ => op2 + | _, None => op1 + | Some p1, Some p2 => Some (fun a => p1 a /\ p2 a) + end. + +Lemma oand1_distr {A}(op1 op2 : OProp1 A) a : + oProp1 (oand1 op1 op2) a <-> oProp1 op1 a /\ oProp1 op2 a. +Proof. + destruct op1 as [ pp1 |], op2 as [ pp2 |]; simpl; + split; try intros [ p1 p2 ]; auto. +Qed. + +Lemma oconj1 {A}(op1 op2 : OProp1 A) a (p1 : oProp1 op1 a)(p2 : oProp1 op2 a) : + oProp1 (oand1 op1 op2) a. +Proof (proj2 (oand1_distr _ _ _) (conj p1 p2)). + +Definition OProp2 A B := option (A -> B -> Prop). +Definition otrue2 {A B} : OProp2 A B := None. +Definition ofalse2 {A B} : OProp2 A B := Some (fun _ _ => False). +Definition oprop2 {A B} : (A -> B -> Prop) -> OProp2 A B := @Some (A -> B -> Prop). + +Definition oProp2 {A B}(op : OProp2 A B) := match op with + | None => fun _ _ => True + | Some p => p + end. + +Definition omap2 {A B A' B'} f (op : OProp2 A B) : OProp2 A' B' := + match op with + | None => None + | Some p => Some (fun a b => f p a b) + end. + +Definition omap21 {A B C} f (op : OProp2 A B) : OProp1 C := + match op with + | None => None + | Some p => Some (fun c => f p c) + end. + +Definition oimply2 {A B}(op1 op2 : OProp2 A B) := match op1, op2 with + | None, _ => op2 + | _, None => None + | Some p1, Some p2 => Some (fun a b => p1 a b -> p2 a b) + end. + +Lemma oimply2_distr {A B}(op1 op2 : OProp2 A B) a b : + oProp2 (oimply2 op1 op2) a b <-> (oProp2 op1 a b -> oProp2 op2 a b). +Proof. + destruct op1 as [ pp1 |], op2 as [ pp2 |]; simpl; split; auto. +Qed. + +Definition oimply2b {A B} p (op : OProp2 A B) := match op with + | None => None + | Some p' => Some (fun a b => p a b -> p' a b) + end. + +Lemma oimply2b_distr {A B} p (op : OProp2 A B) a b : + oProp2 (oimply2b p op) a b <-> (p a b -> oProp2 op a b). +Proof. + destruct op as [ pp |]; simpl; split; auto. +Qed. + +Definition oand2 {A B}(op1 op2 : OProp2 A B) := match op1, op2 with + | None, _ => op2 + | _, None => op1 + | Some p1, Some p2 => Some (fun a b => p1 a b /\ p2 a b) + end. + +Lemma oand2_distr {A B}(op1 op2 : OProp2 A B) a b : + oProp2 (oand2 op1 op2) a b <-> oProp2 op1 a b /\ oProp2 op2 a b. +Proof. + destruct op1 as [ pp1 |], op2 as [ pp2 |]; simpl; + split; try intros [ p1 p2 ]; auto. +Qed. + +Definition oabsorption2 {A B}(op1 op2 : OProp2 A B) := match op1, op2 with + | None, _ => op2 + | _, None => op1 + | Some p1, Some p2 => Some (fun a b => p1 a b /\ (p1 a b -> p2 a b)) + end. + +Lemma oabsorption2_distr {A B}(op1 op2 : OProp2 A B) a b : + oProp2 (oabsorption2 op1 op2) a b <-> + (oProp2 op1 a b /\ (oProp2 op1 a b -> oProp2 op2 a b)). +Proof. + destruct op1 as [ pp1 |], op2 as [ pp2 |]; simpl; + split; try intros [ p1 p2 ]; auto. +Qed. + +Lemma oconj2 {A B}(op1 op2 : OProp2 A B) a b + (p1 : oProp2 op1 a b)(p2 : oProp2 op2 a b) : oProp2 (oand2 op1 op2) a b. +Proof (proj2 (oand2_distr _ _ _ _) (conj p1 p2)). + +Definition olift12a {A B}(op : OProp1 A) : OProp2 A B := + option_map (fun p a b => p a) op. +Definition olift12b {A B}(op : OProp1 B) : OProp2 A B := + option_map (fun p a b => p b) op. + +Definition ounlift21a {A B} a (op : OProp2 A B) : OProp1 B := + option_map (fun p b => p a b) op. +Definition ounlift21b {A B} b (op : OProp2 A B) : OProp1 A := + option_map (fun p a => p a b) op. + +Definition OProp3 A B C := option (A -> B -> C -> Prop). +Definition otrue3 {A B C} : OProp3 A B C := None. +Definition ofalse3 {A B C} : OProp3 A B C := Some (fun _ _ _ => False). +Definition oprop3 {A B C} : (A -> B -> C -> Prop) -> OProp3 A B C + := @Some (A -> B -> C -> Prop). + +Definition oProp3 {A B C}(op : OProp3 A B C) := match op with + | None => fun _ _ _ => True + | Some p => p + end. + +Definition omap3 {A B C A' B' C'} f (op : OProp3 A B C) : OProp3 A' B' C' := + match op with + | None => None + | Some p => Some (fun a b c => f p a b c) + end. + +Definition oand3 {A B C}(op1 op2 : OProp3 A B C) := match op1, op2 with + | None, _ => op2 + | _, None => op1 + | Some p1, Some p2 => Some (fun a b c => p1 a b c /\ p2 a b c) + end. + +Lemma oand3_distr {A B C}(op1 op2 : OProp3 A B C) a b c : + oProp3 (oand3 op1 op2) a b c <-> oProp3 op1 a b c /\ oProp3 op2 a b c. +Proof. + destruct op1 as [ pp1 |], op2 as [ pp2 |]; simpl; + split; try intros [ p1 p2 ]; auto. +Qed. + +Lemma oconj3 {A B C}(op1 op2 : OProp3 A B C) a b c + (p1 : oProp3 op1 a b c)(p2 : oProp3 op2 a b c) : + oProp3 (oand3 op1 op2) a b c. +Proof (proj2 (oand3_distr _ _ _ _ _) (conj p1 p2)). + +Definition olift23ab {A B C}(op : OProp2 A B) : OProp3 A B C := + option_map (fun p a b c => p a b) op. +Definition olift23ac {A B C}(op : OProp2 A C) : OProp3 A B C := + option_map (fun p a b c => p a c) op. +Definition olift23bc {A B C}(op : OProp2 B C) : OProp3 A B C := + option_map (fun p a b c => p b c) op. + +Definition ounlift32a {A B C} a (op : OProp3 A B C) : OProp2 B C := + option_map (fun p b c => p a b c) op. +Definition ounlift32b {A B C} b (op : OProp3 A B C) : OProp2 A C := + option_map (fun p a c => p a b c) op. +Definition ounlift32c {A B C} c (op : OProp3 A B C) : OProp2 A B := + option_map (fun p a b => p a b c) op. + +Lemma OProp1map1 {A A'}(f : _ -> A' -> Prop)(op : OProp1 A) a : + f (fun _ => True) a -> + (oProp1 (omap1 f op) a <-> f (oProp1 op) a). +Proof. + intros fbase; destruct op; simpl; tauto. +Qed. + +(* +Lemma OProp1map2 {A A' B'}(f : _ -> A' -> B' -> Prop)(op : OProp1 A) a b : + f (fun _ => True) a b -> + (oProp2 (option_map f op) a b <-> f (oProp1 op) a b). +Proof. + intros fbase; destruct op; simpl; tauto. +Qed. +*) + +Lemma OProp2map1 {A B A'}(f : _ -> A' -> Prop)(op : OProp2 A B) a : + f (fun _ _ => True) a -> + (oProp1 (omap21 f op) a <-> f (oProp2 op) a). +Proof. + intros fbase; destruct op; simpl; tauto. +Qed. + +Lemma OProp2map2 {A B A' B'}(f : _ -> A' -> B' -> Prop)(op : OProp2 A B) a b : + f (fun _ _ => True) a b -> + (oProp2 (omap2 f op) a b <-> f (oProp2 op) a b). +Proof. + intros fbase; destruct op; simpl; tauto. +Qed. + +Lemma OProp2lift12a {A B} op a b : + @oProp2 A B (olift12a op) a b <-> oProp1 op a. +Proof. + destruct op; simpl; tauto. +Qed. + +Delimit Scope oprop_scope with oprop. +Notation "{{ x }}" := (oprop x%type) + (at level 0, x at level 99) : oprop_scope. +Notation "x /\ y" := (oand x%oprop y%oprop) + (at level 80, right associativity) : oprop_scope. + +Delimit Scope oprop1_scope with oprop1. +Notation "{{ x }}" := (oprop1 x%type) + (at level 0, x at level 99) : oprop1_scope. +Notation "x /\ y" := (oand1 x%oprop1 y%oprop1) + (at level 80, right associativity) : oprop1_scope. +Notation "x m{ f }" := (omap1 (fun p y => p (f y)) x%oprop1) + (at level 25, left associativity) : oprop1_scope. + +Delimit Scope oprop2_scope with oprop2. +Notation "{{ x }}" := (oprop2 x%type) + (at level 0, x at level 99) : oprop2_scope. +Notation "x /\ y" := (oand2 x%oprop2 y%oprop2) + (at level 80, right associativity) : oprop2_scope. +Notation "x m{ f } m{ g }" := (omap2 (fun p y z => p (f y) (g z)) x%oprop2) + (at level 25, left associativity) : oprop2_scope. +Notation "x m{ f } m{ g }" := (omap21 (fun p y => p (f y) (g y)) x%oprop2) + (at level 25, left associativity) : oprop1_scope. + +Delimit Scope oprop3_scope with oprop3. +Notation "{{ x }}" := (oprop3 x%type) + (at level 0, x at level 99) : oprop3_scope. +Notation "x /\ y" := (oand3 x%oprop3 y%oprop3) + (at level 80, right associativity) : oprop3_scope. +Notation "x m{ f } m{ g } m{ h }" := (omap3 (fun p y z w => p (f y) (g z) (h w)) x%oprop3) + (at level 25, left associativity) : oprop3_scope. diff --git a/src/lib/SimpleIndexedMaps.v b/src/lib/SimpleIndexedMaps.v new file mode 100755 index 0000000..0e82de8 --- /dev/null +++ b/src/lib/SimpleIndexedMaps.v @@ -0,0 +1,322 @@ +(* A simpler version of Shu-chun's "indexed map" data structure. + This one just stores things in an association list instead of a tree. *) + +Require Export Coq.Program.Equality. + +Require Import Coqlib. +Require Import Maps. +Require Import DeepSpec.lib.SimpleMaps. + +Record Isomorphism {A} (X Y : A) : Type := { + iso_f : forall F : A -> Type, F X -> F Y; + iso_g : forall F : A -> Type, F Y -> F X + (* := fun F => iso_f (fun x => F x -> F X) id *) +}. + +Arguments iso_f {_ _ _} _ [_] _. +Arguments iso_g {_ _ _} _ [_] _. +Coercion iso_f : Isomorphism >-> Funclass. + +Notation "X ~~~ Y" := (Isomorphism X Y) (at level 70, no associativity). + +Definition identity_automorphism {A X} : @Isomorphism A X X + := {| iso_f F fx := fx; iso_g f fx := fx |}. + +Definition iso_concat {A}{X Y Z : A}(iso_xy : X ~~~ Y)(iso_yz : Y ~~~ Z) := + {| iso_f F fx := iso_yz F (iso_xy F fx); + iso_g F fy := iso_g iso_xy (iso_g iso_yz fy) |}. + +Definition iso_inverse {A}{X Y : A}(iso : X ~~~ Y) := + {| iso_f := iso_g iso; iso_g := iso_f iso |}. + +Notation "p @ q" := (iso_concat p q) (at level 40, left associativity). +Notation "! p" := (iso_inverse p) (at level 35, right associativity). + +Definition isomorphism_compose {A}{X Y Z : A}(iso_xy : X ~~~ Y) + (iso_yz : Y ~~~ Z) : + forall F fx, iso_yz F (iso_xy F fx) = (iso_xy @ iso_yz) F fx. +Proof. + intros F fx; reflexivity. +Qed. + +Lemma isomorphism_decidable_neq_absurd {A}{X Y : A}(dec : A -> bool) : + X ~~~ Y -> dec X = true -> dec Y = false -> False. +Proof. + intros iso dec_x dec_y. + set (F := fun Z : A => match dec Z with + | true => True + | false => False + end). + assert (fx : F X). + { unfold F; rewrite dec_x; exact I. } + + assert (fy : F Y). + { apply iso, fx. } + + unfold F in fy. + rewrite dec_y in fy. + exact fy. +Qed. + +(** Defining the isomorphism using Leibniz equality means that we have to + rely on some properties of functors. One of them, especially, is based + on paramatricity, which can't be proved inside Coq, and is axiomatized + as [functor_identity_paramatricity]. + + Hopefully it is consistent. It implies eq_rect_eq, which is already + implied by proof irrelevance required by CompCert. +*) +Axiom functor_identity_paramatricity : + forall (A : Type)(X : A)(f : forall F, F X -> F X) + (F : A -> Type)(x : F X), x = f F x. + +Module AxiomImplication. + Lemma eq_rect_eq : + forall (U : Type) (p : U) (Q : U -> Type) (x : Q p) (h : p = p), + x = eq_rect p Q x p h. + Proof. + intros. + exact (functor_identity_paramatricity U p + (fun Q x => eq_rect p Q x p h) Q x). + Qed. +End AxiomImplication. + +Lemma isomorphism_identity_gf' {A}{X Y : A}(iso1 iso2 : X ~~~ Y)(F : A -> Type) + (x : F X) : x = iso_g iso1 (iso_f iso2 x). +Proof. + set (iso_comp := iso2 @ !iso1). + change (iso_g iso1 (iso_f iso2 x)) with (iso_f iso_comp x). + apply functor_identity_paramatricity. +Qed. + +Lemma isomorphism_identity_fg' {A}{X Y : A}(iso1 iso2 : X ~~~ Y)(F : A -> Type) + (y : F Y) : y = iso_f iso1 (iso_g iso2 y). +Proof. + set (iso_comp := !iso1 @ iso2). + change (iso_f iso1 (iso_g iso2 y)) with (iso_g iso_comp y). + apply functor_identity_paramatricity. +Qed. + +Lemma isomorphism_identity_gf {A}{X Y : A}(iso : X ~~~ Y)(F : A -> Type) + (x : F X) : x = iso_g iso (iso_f iso x). +Proof (isomorphism_identity_gf' iso iso F x). + +Lemma isomorphism_identity_fg {A}{X Y : A}(iso : X ~~~ Y)(F : A -> Type) + (y : F Y) : y = iso_f iso (iso_g iso y). +Proof (isomorphism_identity_fg' iso iso F y). + +Lemma isomorphism_unique {A}{X Y : A}(iso1 iso2 : X ~~~ Y)(F : A -> Type) + (x : F X) : iso_f iso1 x = iso_f iso2 x. +Proof. + transitivity (iso_f iso2 (iso_g iso2 (iso_f iso1 x))). + apply isomorphism_identity_fg. + f_equal; symmetry. + apply isomorphism_identity_gf'. +Qed. + +Lemma isomorphism_injective {A}{X Y : A}(iso : X ~~~ Y) : + forall F x1 x2, iso F x1 = iso F x2 -> x1 = x2. +Proof. + intros F x1 x2 eq. + rewrite (isomorphism_identity_gf iso F x1). + rewrite eq. + symmetry; apply isomorphism_identity_gf. +Qed. + +(* To avoid useless definitions of inductors in extracted code. *) +Local Unset Elimination Schemes. +Local Unset Case Analysis Schemes. + +Open Scope positive. + +(** * An implementation of indexed trees over type [positive] + Element types are indexed by a tree [PTree.t A] along with a projection + [proj] that extract the actual [Type] from [A]. *) +Module Type TypeProjection. + Variable A : Type. + Variable proj : A -> Type. +End TypeProjection. + +Module IPList(TP : TypeProjection). + Definition elt := positive. + Definition elt_eq := peq. + + Definition OptionalType (a : option TP.A) : Type := + match a with + | None => unit + | Some a' => TP.proj a' + end. + + Inductive ilist : AList.t TP.A -> Type := + | empty : ilist AList.empty + | set : forall {tps : AList.t TP.A} i tp (x : TP.proj tp) (xs : ilist tps), ilist (AList.set i tp tps). + + Definition get_index {idx}(_ : ilist idx) := idx. + + Definition t := ilist. + + + (* Todo: add this to the AList interface. *) + Lemma AList_get_unfold : forall (A:Type) i j (x:A) l, + AList.get i ((j,x)::l) = if (i =? j) then Some x else AList.get i l. + Proof. + reflexivity. + Qed. + + Fixpoint get {idx}(i : positive)(m : t idx){struct m} : + OptionalType (AList.get i idx) := + match m with + | empty => tt + | set tps j tp x xs => + if (i =? j) as b + return (OptionalType (if b then Some tp else AList.get i tps)) + then x + else get i xs + end. + + Definition get_eq {idx}(i : positive){OA}(m : t idx) : + AList.get i idx ~~~ OA -> OptionalType OA + := fun iso => iso OptionalType (get i m). + + (* Fixpoint remove {idx}(i : positive)(m : t idx){struct i} : t (PTree.remove i idx). *) + (* Fixpoint remove_eq {idx}(i : positive) A (m : t (PTree.set i A idx)){struct i} : + PTree.get i idx ~~~ None -> t idx. *) + + (* TODO: move this to the other file. *) + Notation "a ! b" := (AList.get b a) (at level 1). + + Theorem gempty (i : positive)(iso : AList.empty ! i ~~~ None) : + iso OptionalType (get i empty) = tt. + Proof. + simpl. + symmetry. + apply (functor_identity_paramatricity _ _ (iso_f iso) OptionalType). + Qed. + + Lemma gleaf (i : positive)(iso : AList.empty ! i ~~~ None) : + iso OptionalType (get i empty) = tt. + Proof (gempty i iso). + + Theorem get_eq_fusion {idx OA OA'}(i : positive)(m : t idx) + (iso : AList.get i idx ~~~ OA)(iso' : OA ~~~ OA') : + iso' OptionalType (get_eq i m iso) = get_eq i m (iso @ iso'). + Proof. + unfold get_eq. + rewrite isomorphism_compose. + reflexivity. + Qed. + + Theorem gss {idx A}(i : positive)(x : TP.proj A)(m : t idx) + (iso : (AList.set i A idx) ! i ~~~ Some A) : + get i (set i A x m) = iso_g iso (F := OptionalType) x. + Proof. + revert iso. + simpl. + rewrite Pos.eqb_refl. + intros iso. + rewrite <- (functor_identity_paramatricity _ _ (iso_g iso)). + reflexivity. + Qed. + + Theorem gess' {idx A OA}(i : positive)(x : TP.proj A)(m : t idx) + iso (iso' : Some A ~~~ OA) : + get_eq i (OA := OA) (set i A x m) iso = iso' OptionalType x. + Proof. + unfold get_eq. + set (iso_comp := iso @ !iso'). + rewrite (gss _ _ _ iso_comp). + symmetry. + apply (isomorphism_identity_fg iso _ (iso' OptionalType x)). + Qed. + + Theorem gess {idx A}(i : positive)(x : TP.proj A)(m : t idx) iso : + get_eq i (OA := Some A) (set i A x m) iso = x. + Proof (gess' i x m iso identity_automorphism). + + Definition gleaf_iso {i} : (@AList.empty TP.A) ! i ~~~ None. + Proof. destruct i; exact identity_automorphism. Defined. + +Lemma eqb_neq' : forall x y, x<>y -> (x =? y) = false. + Proof. + intros. + rewrite <- Pos.eqb_neq in H. + auto. + Qed. + + Theorem gso {idx A}(i j : positive)(x : TP.proj A)(m : t idx) iso : + i <> j -> get i (set j A x m) = iso_f iso (get i m). + Proof. + intros Hij. + generalize dependent idx. + simpl. + rewrite (eqb_neq' _ _ Hij). + intros. + apply functor_identity_paramatricity. + Qed. + + Theorem geso {idx A OA}(i j : positive)(x : TP.proj A)(m : t idx) iso iso' : + i <> j -> get_eq i (set j A x m) iso = @get_eq _ i OA m iso'. + Proof. + unfold get_eq. + intros neq. + rewrite (gso _ _ x m (iso' @ !iso) neq). + set (iso_comp := (iso' @ !iso) @ iso). + change (iso_f iso (iso_f (iso' @ !iso) (get i m))) with + (iso_f iso_comp (get i m)). + apply isomorphism_unique. + Qed. + + (* + Theorem gres {idx A}(i : positive)(m : t (PTree.set i A idx)) iso : + get i (remove_eq i m iso) = iso_g iso (F := OptionalType) tt. + Proof. + Theorem geres' {idx A OA}(i : positive)(m : t (PTree.set i A idx)) iso iso' : + get_eq i (OA := OA) (remove_eq i m iso) iso' = iso_g (!iso' @ iso) tt. + + Theorem geres {idx A}(i : positive)(m : t (PTree.set i A idx)) iso : + get_eq i (OA := None) (remove_eq i m iso) iso = tt. + + Lemma rleaf {A}(i : positive)(m : t (PTree.set i A PTree.Leaf)) iso : + remove_eq i m iso = Leaf. + + Theorem greo {idx A}(i j : positive)(m : t (PTree.set j A idx)) iso iso' : + i <> j -> get i (remove_eq j m iso) = iso_f iso' (get i m). + + Theorem gereo {idx A OA}(i j : positive)(m : t (PTree.set j A idx)) + iso iso' iso'' : + i <> j -> get_eq i (OA := OA) (remove_eq j m iso) iso' + = iso_f iso' (get_eq i m iso''). + *) +End IPList. + +(* +Module PTreeAux. + Theorem smap1: + forall (A B: Type) (f: A -> B) (a: A) (i: PTree.elt) (m: PTree.t A), + PTree.set i (f a) (PTree.map1 f m) = PTree.map1 f (PTree.set i a m). + Proof. + induction i; destruct m; simpl; try rewrite <- IHi; reflexivity. + Qed. +End PTreeAux. + +Module PMapAux. + Theorem smap: + forall (A B: Type) (f: A -> B) (a: A) (i: PMap.elt) (m: PMap.t A), + PMap.set i (f a) (PMap.map f m) = PMap.map f (PMap.set i a m). + Proof. + intros. + unfold PMap.map, PMap.set; simpl. + f_equal; apply PTreeAux.smap1. + Qed. +End PMapAux. + +Module ZMapAux. + Theorem smap: + forall (A B: Type) (f: A -> B) (a: A) (i: ZMap.elt) (m: ZMap.t A), + ZMap.set i (f a) (ZMap.map f m) = ZMap.map f (ZMap.set i a m). + Proof. + intros. + apply PMapAux.smap. + Qed. +End ZMapAux. +*) diff --git a/src/lib/SimpleMaps.v b/src/lib/SimpleMaps.v new file mode 100755 index 0000000..e6a8194 --- /dev/null +++ b/src/lib/SimpleMaps.v @@ -0,0 +1,386 @@ + +(* Compcert comes with a TREE interface for finite maps, but the Compcert + implementation of it is in terms of balanced trees, which generally does not + reduce well at typechecking-time. + + Here we provide a simple (and slow) implementation of the same + interface in terms just functions. *) + +Require Import Equivalence EquivDec. +Require Import Coqlib. +Require Import Maps. +Require Import Coq.Logic.FunctionalExtensionality. + +(* To avoid useless definitions of inductors in extracted code. *) +Local Unset Elimination Schemes. +Local Unset Case Analysis Schemes. + +Set Implicit Arguments. + + +(* This is a reduced version of Compcert's TREE interface, omitting + things that can't be supported by the simple association-list implementation below. *) +Module Type ALIST. + Variable elt: Type. + Variable elt_eq: forall (a b: elt), {a = b} + {a <> b}. + Variable t: Type -> Type. + Variable empty: forall (A: Type), t A. + Variable get: forall (A: Type), elt -> t A -> option A. + Variable set: forall (A: Type), elt -> A -> t A -> t A. + + Hypothesis gempty: + forall (A: Type) (i: elt), get i (empty A) = None. + Hypothesis gss: + forall (A: Type) (i: elt) (x: A) (m: t A), get i (set i x m) = Some x. + Hypothesis gso: + forall (A: Type) (i j: elt) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Hypothesis gsspec: + forall (A: Type) (i j: elt) (x: A) (m: t A), + get i (set j x m) = if elt_eq i j then Some x else get i m. + (* + Hypothesis gsident: + forall (A: Type) (i: elt) (m: t A) (v: A), + get i m = Some v -> set i v m = m. + *) + + (** Applying a function to all data of a tree. *) + Variable map: + forall (A B: Type), (elt -> A -> B) -> t A -> t B. + Hypothesis gmap: + forall (A B: Type) (f: elt -> A -> B) (i: elt) (m: t A), + get i (map f m) = option_map (f i) (get i m). + + (** Same as [map], but the function does not receive the [elt] argument. *) + Variable map1: + forall (A B: Type), (A -> B) -> t A -> t B. + Hypothesis gmap1: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map1 f m) = option_map f (get i m). + + + (** Enumerating the bindings of a tree. *) + Variable elements: + forall (A: Type), t A -> list (elt * A). + Hypothesis elements_correct: + forall (A: Type) (m: t A) (i: elt) (v: A), + get i m = Some v -> In (i, v) (elements m). + (* + Hypothesis elements_complete: + forall (A: Type) (m: t A) (i: elt) (v: A), + In (i, v) (elements m) -> get i m = Some v. + Hypothesis elements_keys_norepet: + forall (A: Type) (m: t A), + list_norepet (List.map (@fst elt A) (elements m)). + *) +End ALIST. + +Open Scope positive. + + +(* Maps as association lists. + This does not quite fit the interface, because we don't have that extensional + equality implies exact equality, but we don't need that for the DeepSpec development. *) +Module AList <: ALIST. + Definition elt := positive. + Definition elt_eq := peq. + + Definition t (A:Type) := list (positive * A). + + Definition empty {A : Type} : t A := nil. + + + Fixpoint get {A : Type} (i : positive) (m : t A) : option A := + match m with + | (j, a) :: m' => if (i =? j) then Some a else get i m' + | nil => None + end. + + Definition set (A : Type) (i : positive) (v : A) (m : t A) : t A := + (i,v) :: m. + + + (* We don't need remove. + Definition remove (A : Type) (i : positive) (m : t A) {struct i} : t A := + *) + + Theorem gempty: + forall (A: Type) (i: positive), get i empty = (@None A). + Proof. + reflexivity. + Qed. + + Theorem gss: + forall (A: Type) (i: positive) (x: A) (m: t A), get i (set i x m) = Some x. + Proof. + intros. + unfold get,set. + simpl. + rewrite Pos.eqb_refl. + reflexivity. + Qed. + + Theorem gso: + forall (A: Type) (i j: positive) (x: A) (m: t A), + i <> j -> get i (set j x m) = get i m. + Proof. + unfold get, set. + intros. + simpl. + replace (i =? j) with false. + reflexivity. + symmetry; rewrite Pos.eqb_neq; assumption. + Qed. + + Theorem gsspec: + forall (A: Type) (i j: positive) (x: A) (m: t A), + get i (set j x m) = if peq i j then Some x else get i m. + Proof. + intros. + destruct (peq i j); + [ rewrite e; apply gss | apply gso; auto ]. + Qed. + + + (* not true with this representation. *) + (* + Theorem gsident: + forall (A: Type) (i: positive) (m: t A) (v: A), + get i m = Some v -> set i v m = m. + Proof. + induction i; intros; destruct m; simpl; simpl in H; try congruence. + rewrite (IHi m2 v H); congruence. + rewrite (IHi m1 v H); congruence. + Qed. + *) + + (* This is used at one point in the development, in SynthesisStmt.v. + However, that part deals directly with compcert operational semantics, + so we will have had to insert some kind of conversion to PTrees in any case. *) + (* + Theorem set2: + forall (A: Type) (i: elt) (m: t A) (v1 v2: A), + set i v2 (set i v1 m) = set i v2 m. *) + + (* We don't use remove. *) + (* + Theorem grs: + forall (A: Type) (i: positive) (m: t A), get i (remove i m) = None. + Theorem gro: + forall (A: Type) (i j: positive) (m: t A), + i <> j -> get i (remove j m) = get i m. + Theorem grspec: + forall (A: Type) (i j: elt) (m: t A), + get i (remove j m) = if elt_eq i j then None else get i m. + *) + + (* We could prove these, but they are not used in the development. *) + (* + Section BOOLEAN_EQUALITY. + Variable A: Type. + Variable beqA: A -> A -> bool. + + Fixpoint bempty (m: t A) : bool := + + Fixpoint beq (m1 m2: t A) {struct m1} : bool := + + Lemma bempty_correct: + forall m, bempty m = true <-> (forall x, get x m = None). + + Lemma beq_correct: + forall m1 m2, + beq m1 m2 = true <-> + (forall (x: elt), + match get x m1, get x m2 with + | None, None => True + | Some y1, Some y2 => beqA y1 y2 = true + | _, _ => False + end). + End BOOLEAN_EQUALITY. + *) + + + + Definition map (A B : Type) (f : positive -> A -> B) m := List.map (fun p => (fst p, f (fst p) (snd p))) m. + + Theorem gmap: + forall (A B: Type) (f: positive -> A -> B) (i: positive) (m: t A), + get i (map f m) = option_map (f i) (get i m). + Proof. + induction m. + + reflexivity. + + destruct a as [j a]. + unfold get. + simpl. + destruct (i =? j) eqn:Heq. + * rewrite Peqb_eq in Heq. + subst. + reflexivity. + * unfold get in IHm. + rewrite IHm. + reflexivity. + Qed. + + Definition map1 (A B: Type) (f: A -> B) (m: t A) : t B := + List.map (fun p => (fst p, f (snd p))) m. + + Theorem gmap1: + forall (A B: Type) (f: A -> B) (i: elt) (m: t A), + get i (map1 f m) = option_map f (get i m). + Proof. + induction m. + + reflexivity. + + destruct a as [j a]. + simpl. + unfold get. + simpl. + destruct (i =? j)%positive eqn:Heq. + * rewrite Peqb_eq in Heq. + subst. + reflexivity. + * unfold get in IHm. + rewrite IHm. + reflexivity. + Qed. + + (* not used. + Fixpoint filter1 (A: Type) (pred: A -> bool) (m: t A) {struct m} : t A := + + Theorem gfilter1: + forall (A: Type) (pred: A -> bool) (i: elt) (m: t A), + get i (filter1 pred m) = + match get i m with None => None | Some x => if pred x then Some x else None end. + Proof. + *) + + Definition elements (A: Type) (m: t A) := m. + + Theorem elements_correct: + forall (A: Type) (m: t A) (i: positive) (v: A), + get i m = Some v -> In (i, v) (elements m). + Proof. + induction m. + + compute. + intros; congruence. + + destruct a as [j a]. + simpl. + unfold get. + simpl. + intros. + destruct (i =? j)%positive eqn:Heq. + * rewrite Peqb_eq in Heq. + subst. + left. + congruence. + * unfold get in IHm. + specialize IHm with i v. + rewrite H in IHm. + right; auto. + Qed. + + (* This *was* used, a lot, but it's not true using this representation. *) + (* + Theorem elements_complete: + forall (A: Type) (m: t A) (i: positive) (v: A), + In (i, v) (elements m) -> get i m = Some v. + Proof. + *) + + (* Happily, the following is good enough for our purposes. *) + Theorem elements_domain_complete: + forall (A: Type) (m: t A) (i: positive), + In i (List.map fst (elements m)) -> exists v, get i m = Some v. + Proof. + intros A m i H. + induction m. + - destruct H. + - idtac. + destruct a as [j u]. + destruct (Pos.eq_dec j i)%positive as [Heq | Hneq]. + + exists u. + rewrite Heq. + unfold AList.get. + rewrite Pos.eqb_refl. + reflexivity. + + idtac. + change ((j,u)::m) with (AList.set j u m). + rewrite AList.gso by auto. + apply IHm. + destruct H. + * simpl in H. + tauto. + * exact H. + Qed. + + + Lemma elements_set_monotone {A} x i a (m : AList.t A) : + In x (List.map fst (elements m)) -> + In x (List.map fst (elements (set i a m))). + Proof. + destruct (peq x i) as [ i_eq | i_ne ]; subst. + - unfold AList.set, AList.elements. + simpl. + auto. + - unfold AList.set, AList.elements. + simpl; right; auto. + Qed. + + (* Ditto. *) + (* + Theorem elements_keys_norepet: + forall (A: Type) (m: t A), + list_norepet (List.map (@fst elt A) (elements m)). + + Theorem elements_canonical_order: + forall (A B: Type) (R: A -> B -> Prop) (m: t A) (n: t B), + (forall i x, get i m = Some x -> exists y, get i n = Some y /\ R x y) -> + (forall i y, get i n = Some y -> exists x, get i m = Some x /\ R x y) -> + list_forall2 + (fun i_x i_y => fst i_x = fst i_y /\ R (snd i_x) (snd i_y)) + (elements m) (elements n). + *) + + (* Not true, not used. *) + (* + Theorem elements_extensional: + forall (A: Type) (m n: t A), + (forall i, get i m = get i n) -> + elements m = elements n. + *) + + (* + Definition fold (A B : Type) (f: B -> positive -> A -> B) (m: t A) (v: B) := + xfold f xH m v. + + Theorem fold_spec: + forall (A B: Type) (f: B -> positive -> A -> B) (v: B) (m: t A), + fold f m v = + List.fold_left (fun a p => f a (fst p) (snd p)) (elements m) v. + *) + + (* + Fixpoint fold1 (A B: Type) (f: B -> A -> B) (m: t A) (v: B) {struct m} : B := + + Theorem fold1_spec: + forall (A B: Type) (f: B -> A -> B) (v: B) (m: t A), + fold1 f m v = + List.fold_left (fun a p => f a (snd p)) (elements m) v. + Proof. + *) + +End AList. + + +Definition dummy : nat := (42%nat). + +Fixpoint PTree_of_AList {A:Type} (m: AList.t A) : PTree.t A := + match m with + | nil => PTree.empty A + | (i,a)::m => PTree.set i a (PTree_of_AList m) + end. + + +Close Scope positive. + +Notation "a ! b" := (AList.get b a) (at level 1) : alist_scope. +Delimit Scope alist_scope with alist. diff --git a/src/minic/dune b/src/minic/dune new file mode 100644 index 0000000..9e492c8 --- /dev/null +++ b/src/minic/dune @@ -0,0 +1,6 @@ +(executable + (name minicc) + (modes byte exe) + (libraries minic) +; (preprocess (action (run %{bin:cppo} -V OCAML:%{ocaml_version} %{input-file}))) +) diff --git a/src/minic/minicc.ml b/src/minic/minicc.ml new file mode 100644 index 0000000..0f0af0d --- /dev/null +++ b/src/minic/minicc.ml @@ -0,0 +1,56 @@ +open Minic +open Backend.BytecodeExt + +let usage () = print_endline + "usage: minicc file (print(-verbose)? | assembly(-runtime)? | (bytecode | ewasm)(-runtime)?)?"; + ignore (exit 1); + +type modes = PRINT | VERBOSE | ASSEMBLY | ASSEMBLY_RUNTIME | BYTECODE | BYTECODE_RUNTIME | EWASM | EWASM_RUNTIME + +let mode_of_string = function + | "print" -> PRINT + | "print-verbose" -> VERBOSE + | "assembly" -> ASSEMBLY + | "assembly-runtime" -> ASSEMBLY_RUNTIME + | "bytecode" -> BYTECODE + | "bytecode-runtime" -> BYTECODE_RUNTIME + | "ewasm" -> EWASM + | "ewasm-runtime" -> EWASM_RUNTIME + | _ -> usage (); exit 1 + +let minic mode filename = + let chan = open_in filename in + let lexbuf = Lexing.from_channel chan in + let parsed = try + Parser.file Lexer.token lexbuf + with Failure _ | Parser.Error -> + let curr = lexbuf.Lexing.lex_curr_p in + let line = curr.Lexing.pos_lnum in + let cnum = curr.Lexing.pos_cnum - curr.Lexing.pos_bol in + let tok = Lexing.lexeme lexbuf in + print_endline (String.concat "" + [ filename; ":"; string_of_int line; ":"; string_of_int cnum; + ": Syntax error at token \""; tok; "\"." ]); + exit 1 + in + let name_tbls, genv = Typecheck.typecheck parsed in + match mode with + | PRINT -> + print_endline (Backend.LanguageExt.show_genv false name_tbls genv) + | VERBOSE -> + print_endline (Backend.LanguageExt.show_genv true name_tbls genv) + | ASSEMBLY -> print_endline (assembly false genv) + | ASSEMBLY_RUNTIME -> print_endline (assembly true genv) + | BYTECODE -> print_endline (bytecode false genv) + | BYTECODE_RUNTIME -> print_endline (bytecode true genv) + | EWASM -> print_endline (ewasm false genv) + | EWASM_RUNTIME -> print_endline (ewasm true genv) + +let main () = + if (Array.length Sys.argv) = 1 then usage (); + let mode = + if (Array.length Sys.argv) >= 3 then mode_of_string Sys.argv.(2) + else PRINT in + minic mode Sys.argv.(1) + +let () = main () diff --git a/src/minic/src/ast.ml b/src/minic/src/ast.ml new file mode 100644 index 0000000..ec6365e --- /dev/null +++ b/src/minic/src/ast.ml @@ -0,0 +1,58 @@ +open Backend + +type p_type = + | Pvoid + | Pint of string + | Pbool + | Ppointer of p_type + | Parray of p_type * int + | Pmapping of p_type * p_type + | Pstruct of string * p_variable_declaration list + | Punion of string * p_variable_declaration list + | Puserstruct of string +and p_variable_declaration = bool * p_type * string + +type p_expr = + (* Econst_int or Econst_int256 *) + | Pconst of string + (* Eglob, Evar and Etempvar *) + | Pvar of string + | Pderef of p_expr + | Punop of Cop.unary_operation * p_expr + | Pbinop of p_expr * Cop.binary_operation * p_expr + | Pfield of p_expr * string + (* Earrayderef and Ehashderef *) + | Parrayderef of p_expr * p_expr + | Pcall0 of MachineModel.builtin0 + | Pcall1 of MachineModel.builtin1 * p_expr + +type p_statement = + (* Sassign, Sset, Stransfer *) + | Passign of p_expr * p_expr + | Pifthenelse of p_expr * p_statements * p_statements option + | Ploop of p_statements + | Pbreak + | Preturn of p_expr option + | Ptransfer of p_expr * p_expr + | Prevert + | Pcall of string option * string * p_expr list + | Plog of p_expr list * p_expr list + | Pcallmethod of p_expr * p_expr list * p_expr * p_expr * p_expr option * p_expr list +and p_statements = p_statement list + + +type visibility = Public | Private + +type p_function = { p_fn_visibility: visibility option; + p_fn_name: string; + p_fn_return : p_type; + p_fn_params : p_variable_declaration list; + p_fn_temps : p_variable_declaration list; + p_fn_body : p_statement list } + +type p_declaration = + | Ptype_decl of p_type + | Pvar_decl of p_variable_declaration + | Pfunc of p_function + +type p_file = p_declaration list diff --git a/src/minic/src/astExt.ml b/src/minic/src/astExt.ml new file mode 100644 index 0000000..024f9ba --- /dev/null +++ b/src/minic/src/astExt.ml @@ -0,0 +1,16 @@ +open Ast +open Backend + open CopExt + open Printf + open MachineModelExt + +let rec show_p_expr = function + | Pconst s -> s + | Pvar s -> s + | Pderef e -> "*" ^ (show_p_expr e) + | Punop (unop, e) -> (show_unop_symb unop) ^ (show_p_expr e) + | Pbinop (e1, binop, e2) -> sprintf "(%s %s %s)" (show_p_expr e1) (show_binop_symb binop) (show_p_expr e2) + | Pfield (e, s) -> sprintf "%s.%s" (show_p_expr e) s + | Parrayderef (e1, e2) -> sprintf "%s[%s]" (show_p_expr e1) (show_p_expr e2) + | Pcall0 b -> show_builtin0 b + | Pcall1 (b, e) -> sprintf "%s(%s)" (show_builtin1 b) (show_p_expr e) diff --git a/src/minic/src/dune b/src/minic/src/dune new file mode 100644 index 0000000..7a95a77 --- /dev/null +++ b/src/minic/src/dune @@ -0,0 +1,11 @@ +(library + (name minic) + (libraries backend core) +) + +(ocamllex lexer) + +(menhir + (modules parser) + (flags --unused-precedence-levels) +) diff --git a/src/minic/src/lexer.mll b/src/minic/src/lexer.mll new file mode 100644 index 0000000..33b513b --- /dev/null +++ b/src/minic/src/lexer.mll @@ -0,0 +1,93 @@ +{ +open Parser +open Backend.AST + +let trim s n = + String.sub s n ((String.length s) - n) + +} + +let letter = ['A'-'Z' 'a'-'z' '_'] +let digit = ['0'-'9'] +let whitespace = [' ' '\t' '\r'] +let decimal = ['1'-'9'] digit* | '0' + +rule token = parse + | "\n" { Lexing.new_line lexbuf; token lexbuf } + | whitespace+ { token lexbuf } + | decimal as d { NUM d } + | ";" { SEMICOLON } + | "," { COMMA } + | "[" { LBRACKET } + | "]" { RBRACKET } + | "{" { LBRACE } + | "}" { RBRACE } + | "(" { LPAREN } + | ")" { RPAREN } + | "#" { HASH } + | "=" { EQUAL } + | "<-" { LARROW } + | "//" { comment lexbuf } + | "/*" { multi_comment 0 lexbuf } + | "+" { ADD } + | "-" { SUB } + | "*" { MUL } + | "/" { DIV } + | "%" { MOD } + | "**" { EXP } + | "&" { AND } + | "|" { OR } + | "^" { XOR } + | "<<" { SHL } + | ">>" { SHR } + | "==" { EQ } + | "!=" { NE } + | "<" { LT } + | ">" { GT } + | "<=" { LE } + | ">=" { GE } + | "!" { NOTBOOL } + | "~" { NOTINT } + | "." { DOT } + | "sha1" { SHA1 } + | "sha2" { SHA2 } + | "if" { IF } + | "else" { ELSE } + | "void" { VOID } + | "int" decimal? as s { INT (trim s 3) } + | "bool" { BOOL } + | "public" { PUBLIC } + | "private" { PRIVATE } + | "while" { WHILE } + | "break" { BREAK } + | "return" { RETURN } + | "address" { ADDRESS } + | "origin" { ORIGIN } + | "caller" { CALLER } + | "callvalue" { CALLVALUE } + | "coinbase" { COINBASE } + | "timestamp" { TIMESTAMP } + | "number" { NUMBER } + | "balance" { BALANCE } + | "blockhash" { BLOCKHASH } + | "transfer" { TRANSFER } + | "revert" { REVERT } + | "struct" { STRUCT } + | "union" { UNION } + | "emit" { EMIT } + | "callmethod" { CALLMETHOD } + | "memory" { MEMORY } + | "chainid" { CHAINID } + | "selfbalance" { SELFBALANCE } + | letter (letter | digit)* as s { IDENT s } + | eof { EOF } + +and comment = parse + | '\n' { Lexing.new_line lexbuf; token lexbuf } + | _ { comment lexbuf } + +and multi_comment count = parse + | "/*" { multi_comment (count+1) lexbuf } + | "*/" { if count > 0 then multi_comment (count-1) lexbuf else token lexbuf } + | '\n' { Lexing.new_line lexbuf; multi_comment count lexbuf } + | _ { multi_comment count lexbuf } diff --git a/src/minic/src/parser.mly b/src/minic/src/parser.mly new file mode 100644 index 0000000..3f1ce4a --- /dev/null +++ b/src/minic/src/parser.mly @@ -0,0 +1,249 @@ +%{ +open Ast +open Printf + +open Backend + open AST + open BinNumsExt + open Ctypes + open Cop + module D = Datatypes + open Integers + open Language + +%} + +%token NUM +%token SEMICOLON +%token COMMA +%token LBRACKET +%token RBRACKET +%token LBRACE +%token RBRACE +%token LPAREN +%token RPAREN +%token HASH +%token EQUAL +%token LARROW + +%token ADD +%token SUB +%token MUL +%token DIV +%token MOD +%token EXP +%token AND +%token OR +%token XOR +%token SHL +%token SHR +%token EQ +%token NE +%token LT +%token GT +%token LE +%token GE + +%token NOTBOOL +%token NOTINT + +%token DOT + +%token IF +%token ELSE +%token VOID +%token INT +%token BOOL +%token PUBLIC +%token PRIVATE +%token WHILE +%token BREAK +%token RETURN + +%token ADDRESS +%token ORIGIN +%token CALLER +%token CALLVALUE +%token COINBASE +%token TIMESTAMP +%token NUMBER + +%token SHA1 +%token SHA2 + +%token BALANCE +%token BLOCKHASH + +%token TRANSFER +%token REVERT + +%token UNION +%token STRUCT + +%token EMIT +%token CALLMETHOD +%token MEMORY + +%token IDENT +%token EOF + +%token CHAINID +%token SELFBALANCE + +%token STRUCT_DEF +%token STRUCT_INSTANCE +%nonassoc STRUCT_INSTANCE +%nonassoc STRUCT_DEF + +%nonassoc EQ NE LT LE GT GE +%left SHA2 +%left OR +%left XOR +%left AND +%left MOD +%left SHL SHR +%left ADD SUB +%left MUL DIV +%left EXP +%right NOTBOOL + +%start file + +%% + +file: + d=declaration* EOF { d } + +declaration: + | t = type_declaration { Ptype_decl t } + | v = variable_declaration { Pvar_decl v } + | f = function_definition { Pfunc f } + +type_declaration: + | t=struct_type_expression SEMICOLON { t } + +variable_declaration: + | v=variable SEMICOLON { v } + +function_definition: + | v=ioption(visibility) t=type_expression i=IDENT + LPAREN p=separated_list(COMMA, variable) RPAREN + LBRACE tmps=variable_declaration* body=statement* RBRACE { + { + p_fn_visibility = v; + p_fn_name = i; + p_fn_return = t; + p_fn_params = p; + p_fn_temps = tmps; + p_fn_body = body; + }} + +visibility: + | PUBLIC { Public } + | PRIVATE { Private } + +statement: + | s=semicolon_statement SEMICOLON { s } + | LBRACE s=statement RBRACE { s } + | IF LPAREN e=expression RPAREN s1=then_clause s2=else_clause? { Pifthenelse (e, s1, s2) } + | WHILE s=then_clause { Ploop s } + + +then_clause: + | s=statement { [s] } + | LBRACE s=statement* RBRACE { s } + +else_clause: + | ELSE s=then_clause { s } + +semicolon_statement: + | e1=expression EQUAL e2=expression { Passign (e1, e2) } + | BREAK { Pbreak } + | RETURN e=expression? { Preturn e } + | TRANSFER LPAREN e1=expression COMMA e2=expression RPAREN { Ptransfer (e1, e2) } + | REVERT { Prevert } + | EMIT LPAREN e1=separated_list(COMMA, expression) SEMICOLON e2=separated_list(COMMA, expression) RPAREN { Plog (e1, e2) } + | i=ioption(terminated(IDENT, LARROW)) func=IDENT LPAREN e=separated_list(COMMA, expression) RPAREN + { Pcall (i, func, e) } + | CALLMETHOD LPAREN + addr=expression SEMICOLON + ret=separated_list(COMMA, expression) SEMICOLON + signature=expression SEMICOLON + value=expression SEMICOLON + gas=expression? SEMICOLON + args=separated_list(COMMA, expression) + RPAREN + { Pcallmethod (addr, ret, signature, value, gas, args) } + +expression: + | LPAREN e=expression RPAREN { e } + | i=IDENT { Pvar i } + | n=NUM { Pconst n } + | MUL e=expression { Pderef e } %prec NOTBOOL + | u=unop e=expression { Punop (u, e) } %prec NOTBOOL + | e1=expression b=binop e2=expression { Pbinop (e1, b, e2) } + | e1=expression LBRACKET e2=expression RBRACKET { Parrayderef (e1, e2) } + | b=builtin0 { Pcall0 b } + | b=builtin1 LPAREN e=expression RPAREN { Pcall1 (b, e) } + | e=expression DOT i=IDENT { Pfield (e, i) } + +%inline binop: + | ADD { Oadd } + | SUB { Osub } + | MUL { Omul } + | DIV { Odiv } + | MOD { Omod } + | EXP { Oexp } + | AND { Oand } + | OR { Oor } + | XOR { Oxor } + | SHL { Oshl } + | SHR { Oshr } + | EQ { Oeq } + | NE { One } + | LT { Olt } + | GT { Ogt } + | LE { Ole } + | GE { Oge } + | SHA2 { Osha_2 } + +type_expression: + | VOID { Pvoid } + | BOOL { Pbool } + | i=INT { Pint i } + | t=type_expression LBRACKET n=NUM RBRACKET { Parray (t, int_of_string n) } + | t1=type_expression HASH t2=delimited(LBRACKET, type_expression, RBRACKET) + { Pmapping (t1, t2) } + | t=type_expression MUL { Ppointer t } + | t=struct_type_expression { t } %prec STRUCT_DEF + | STRUCT i=IDENT { Puserstruct i } %prec STRUCT_INSTANCE + | UNION i=IDENT { Puserstruct i } %prec STRUCT_INSTANCE + + +struct_type_expression: + | STRUCT i=IDENT LBRACE v=variable_declaration* RBRACE { Pstruct (i, v) } + | UNION i=IDENT LBRACE v=variable_declaration* RBRACE { Punion (i, v) } + +variable: + | m=ioption(MEMORY) t=type_expression i=IDENT { Option.is_some m, t, i } + +unop: + | NOTBOOL { Onotbool } + | NOTINT { Onotint } + | SUB { Oneg } + | SHA1 { Osha_1 } + +builtin0: + | ADDRESS { Baddress } + | ORIGIN { Borigin } + | CALLER { Bcaller } + | CALLVALUE { Bcallvalue } + | COINBASE { Bcoinbase } + | TIMESTAMP { Btimestamp } + | NUMBER { Bnumber } + | CHAINID { Bchainid } + | SELFBALANCE { Bselfbalance } + +builtin1: + | BALANCE { Bbalance } + | BLOCKHASH { Bblockhash } diff --git a/src/minic/src/typecheck.ml b/src/minic/src/typecheck.ml new file mode 100644 index 0000000..423589f --- /dev/null +++ b/src/minic/src/typecheck.ml @@ -0,0 +1,351 @@ +(* TODO: Change var store to glob store *) + +open Ast +open AstExt +open Printf + +open Backend + open Ctypes + open CtypesExt + open BinNums + open BinNumsExt + module D = Datatypes + open DatatypesExt + open Globalenvs + open Language + open LanguageExt + open ExpMiniC + open StmtMiniC + +let raise_msg s = Core.raise_s (Core.Sexp.Atom s) + +let match_intsize = function + | "8" -> I8 + | "16" -> I16 + | "32" -> I32 + | "256" | "" -> I256 + | _ as s -> raise_msg (sprintf "invalid intsize: %s" s) + +type ('a, 'b) ident_store = { + tbl: ('a, 'b) Hashtbl.t; + mutable lst: 'a List.t +} + +let glob_store = { tbl = Hashtbl.create 0; lst = [] } +let func_store = { tbl = Hashtbl.create 0; lst = [] } +let method_store = { tbl = Hashtbl.create 0; lst = [] } + +let constructor = ref None + +let struct_field_tbl = Hashtbl.create 0 +let struct_type_tbl = Hashtbl.create 0 +let struct_name_tbl = Hashtbl.create 0 + +let name_tbls = NameTablesExt.empty_name_tables + +let flip_tbl tbl = + let new_tbl = Hashtbl.create (Hashtbl.length tbl) in + Hashtbl.iter (fun name (ident, _) -> Hashtbl.add new_tbl ident name) tbl; + new_tbl + +let strip_prefix prefix name = + let prefix_len = String.length prefix in + let name_len = String.length name in + if name_len > prefix_len && + String.sub name 0 prefix_len = prefix + then + Some (int_of_string (String.sub name prefix_len (name_len - prefix_len))) + else + None + +let rec match_type = function + | Pvoid -> Tvoid + | Pint i -> Tint (match_intsize i, Unsigned) + | Pbool -> Tint (IBool, Unsigned) + | Ppointer t -> Tpointer (Coq_mem, (match_type t)) (* Coq_mem is a dummy val *) + | Parray (t, n) -> Tarray (match_type t, Zpos (positive_of_int n)) + | Pmapping (t1, t2) -> Thashmap (match_type t1, match_type t2) + | Pstruct (i, v) -> let (i', f) = add_struct i v in + let t = Tstruct (i', f) in Hashtbl.add struct_type_tbl i t; t + | Punion (i, v) -> let (i', f) = add_struct i v in + let t = Tunion (i', f) in Hashtbl.add struct_type_tbl i t; t + | Puserstruct s -> Hashtbl.find struct_type_tbl s +(* add a struct or union *) +and add_struct i v = + (* let _, t, i = v in *) + (* can try: store by ident, and have an ident->name table for checking duplicates *) + (* if Hashtbl.mem struct_field_tbl i then *) + (* raise_msg (sprintf "redeclaration of struct type %s" i); *) + let field_tbl = Hashtbl.create 0 in + List.iter (fun (_, t, i') -> + let field_num = (match strip_prefix "field_" i' with + | Some n -> positive_of_int n + | None -> positive_of_int ((Hashtbl.length field_tbl) + 1) + ) in + Hashtbl.add field_tbl i' (field_num, match_type t)) v; + let struct_num = positive_of_int ((Hashtbl.length struct_name_tbl)+1) in + Hashtbl.add struct_field_tbl i field_tbl; + Hashtbl.add struct_name_tbl struct_num i; + let field_list = Hashtbl.fold (fun _ (n, t) prev -> Fcons (n, t, prev)) field_tbl Fnil in + struct_num, field_list + +let process_var store prefix var = + let m, t, i = var in + if Hashtbl.mem store.tbl i then + raise_msg (sprintf "redeclaration of variable %s" i); + let i' = match strip_prefix prefix i with + | Some i -> positive_of_int i + | None -> positive_of_int ((Hashtbl.length store.tbl) + 1) in + let t' = match_type t in + Hashtbl.add store.tbl i (i', (t', m)); + store.lst <- store.lst @ [i] + +let process_glob store var = + let m, t, i = var in + if Hashtbl.mem store.tbl i then + raise_msg (sprintf "redeclaration of variable %s" i); + let i' = match strip_prefix "global_" i with + | Some i -> positive_of_int i + | None -> positive_of_int ((Hashtbl.length store.tbl) + 1) in + let t' = match_type t in + Hashtbl.add store.tbl i (i', t'); + store.lst <- store.lst @ [i] + +let process_param store prefix var = + let () = process_var store prefix var in + let _, _, i = var in let i', (t', m) = Hashtbl.find store.tbl i in + D.Coq_pair (i', t') + +let process_temp store prefix var = + let () = process_var store prefix var in + let _, _, i = var in let i', (t', m) = Hashtbl.find store.tbl i in + m, D.Coq_pair (i', t') + +let func_defined name = + if (Hashtbl.mem func_store.tbl name) + || (Hashtbl.mem method_store.tbl name) + || (name = "constructor" && (Option.is_some !constructor)) then + raise_msg (sprintf "redefinition of function %s" name) + +let add_func name = + let () = func_defined name in + let i = match strip_prefix "func_" name with + | Some n -> positive_of_int n + | None -> positive_of_int ((Hashtbl.length func_store.tbl) + 1) in + Hashtbl.add func_store.tbl name (i, empty_coq_function); + func_store.lst <- func_store.lst @ [name] + +let replace_func name f = + let (i, _) = Hashtbl.find func_store.tbl name in + Hashtbl.replace func_store.tbl name (i, f); + i + +let add_method name = + let () = func_defined name in + let i = match strip_prefix "method_" name with + | Some n -> coq_Z_of_int n + | None -> coq_Z_of_int ((Hashtbl.length method_store.tbl) + 1) in + Hashtbl.add method_store.tbl name (i, empty_coq_function); + method_store.lst <- method_store.lst @ [name] + +let replace_method name f = + let (i, _) = Hashtbl.find method_store.tbl name in + Hashtbl.add method_store.tbl name (i, f); + i + +let rec match_expression tmp_tbl e = + let match_ex = match_expression tmp_tbl in + match e with + (* currently setting to the default int: unsigned 256 *) + | Pconst s -> + Econst_int256 (coq_Z_of_Z (Z.of_string s), match_type (Pint "")) + | Pvar s -> + if Hashtbl.mem tmp_tbl s then + let i, (t, m) = Hashtbl.find tmp_tbl s in + (match m with + | true -> Evar (i, t) + | false -> Etempvar (i, t)) + else if Hashtbl.mem glob_store.tbl s then + let i, t = Hashtbl.find glob_store.tbl s in + Eglob (i, t) + else + raise_msg (sprintf "undefined variable %s" s) + | Pderef e -> (match e with + | Pvar _ -> let e' = match_ex e in Ederef (e', typeof e') + | _ -> raise_msg (sprintf "cannot deref non-variable %s" (show_p_expr e)) ) + | Punop (u, e) -> let e' = match_ex e in Eunop (u, e', typeof e') + | Pbinop (e1, b, e2) -> + let e1' = match_ex e1 in + let e2' = match_ex e2 in + if ((typeof e1') <> (typeof e2')) then + raise_msg (sprintf "types of %s and %s incompatible: %s and %s" + (show_p_expr e1) (show_p_expr e2) + (show_type (typeof e1')) (show_type (typeof e2'))); + Ebinop (b, e1', e2', typeof e1') + | Pfield (e, i) -> + let e' = match_ex e in + let n = match typeof e' with + | Tstruct (i, _) -> i + | Tunion (i, _) -> i + | _ -> raise_msg (sprintf "not a struct: %s" (show_p_expr e)) in + (* Hashtbl.iter (fun a b -> print_endline ((show_pos a) ^ b)) struct_name_tbl; *) + let s = Hashtbl.find struct_name_tbl n in + let n', t = Hashtbl.find (Hashtbl.find struct_field_tbl s) i in + Efield (e', n', t) + | Parrayderef (e1, e2) -> let e1' = match_ex e1 in ( + match typeof e1' with + | Tarray (t, _) -> Eindex (e1', match_ex e2, t) + | Thashmap (_, t) -> Eindex (e1', match_ex e2, t) + | _ -> raise_msg (sprintf "%s: not an array or hashmap" (show_p_expr e1))) + | Pcall0 b -> Ecall0 (b, match_type (Pint "256")) + | Pcall1 (b, e) -> Ecall1 (b, match_ex e, match_type (Pint "256")) + + +let matach_ex_opt tmp_tbl = function + | None -> D.None + | Some e -> D.Some (match_expression tmp_tbl e) + +let rec flatten_statements match_func = function + | [] -> Sskip + | hd::tl -> Ssequence (match_func hd, flatten_statements match_func tl) + +let rec match_statement tmp_tbl in_loop stmt = + let match_ex = match_expression tmp_tbl in + let match_ex_list e = coqlist_of_list (List.map match_ex e) in + let loop = match stmt with Ploop _ -> true | _ -> false in + let in_loop' = loop || in_loop in + let flat_stmts = flatten_statements (match_statement tmp_tbl in_loop') in + match stmt with + | Passign (e1, e2) -> ( + let e1' = match_ex e1 in + let e2' = match_ex e2 in + match e1' with + | Etempvar (i, _) -> Sset (i, e2') + | Evar _ | Eglob _ | Eindex _ | Efield _ -> Sassign (e1', e2') + | _ -> raise_msg (sprintf "cannot assign to nonvariable %s" (show_p_expr e1) + ) + ) + | Pifthenelse (e, s1, s2) -> + let s2' = match s2 with + | None -> Sskip + | Some s -> flat_stmts s in + Sifthenelse (match_ex e, flat_stmts s1, s2') + | Ploop s -> Sloop (flat_stmts s) + | Pbreak -> if not in_loop then raise_msg ("break statment must be in a loop"); + Sbreak + | Preturn e -> (match e with + | None -> Sreturn None + | Some e' -> (match match_ex e' with + | Etempvar (i, _) -> Sreturn (Some i) + | _ -> raise_msg (sprintf "not a tempvar: %s" (show_p_expr e')) + ) + ) + | Ptransfer (e1, e2) -> Stransfer (match_ex e1, match_ex e2) + | Prevert -> Srevert + | Pcall (i, f, e) -> + let i' = match i with + | None -> D.None + | Some s -> D.Some (fst (Hashtbl.find tmp_tbl s)) in + let f' = fst (Hashtbl.find func_store.tbl f) in + let e' = coqlist_of_list (List.map match_ex e) in + Scall(i', f', e') + | Plog (e1, e2) -> Slog (match_ex_list e1, match_ex_list e2) + | Pcallmethod (addr, ret, signature, value, gas, args) -> + let ret' = coqlist_of_list (List.map (fun e -> match match_ex e with + | Etempvar (i, _) -> i + | _ -> raise_msg (sprintf "Not a temp variable: %s" (show_p_expr e)) + ) + ret) in + let signature' = match match_ex signature with + | Econst_int256 (i, _) -> i + | _ -> raise_msg (sprintf "Not an int: %s" (show_p_expr signature)) + in + Scallmethod (match_ex addr, ret', signature', match_ex value, + matach_ex_opt tmp_tbl gas, match_ex_list args) + +let preprocess_func f = + match f.p_fn_name with + | "constructor" -> () + | _ -> match f.p_fn_visibility with + | None -> add_method f.p_fn_name + (* default to public *) + | Some v -> match v with + | Private -> add_func f.p_fn_name + | Public -> add_method f.p_fn_name + +let append_return body = + let rev = List.rev body in + let appended = match rev with + | [] -> [ Sreturn None ] + | hd::_ -> match hd with + | Sreturn _ -> rev + | _ -> (Sreturn None) :: rev + in + List.rev appended + +let process_func f = + let tmp_store = { tbl = Hashtbl.create 0; lst = [] } in + let params = List.map (process_param tmp_store "tmp_") f.p_fn_params in + let vars = List.map (process_temp tmp_store "tmp_") f.p_fn_temps in + let filter flip (m, it) = match flip m with + | true -> Some it + | false -> None in + let temps = List.filter_map (filter (fun b -> not b)) vars in + let locals = List.filter_map (filter (fun b -> b)) vars in + let body = List.map (match_statement tmp_store.tbl false) f.p_fn_body in + let body' = append_return body in + let f' = { fn_return = (match_type f.p_fn_return); + fn_params = coqlist_of_list params; + fn_temps = coqlist_of_list temps; + fn_locals = coqlist_of_list locals; + fn_body = flatten_statements (fun x -> x) body'; + } in + match f.p_fn_name with + | "constructor" -> func_defined f.p_fn_name; + constructor := Some f'; + name_tbls.constructor_tmps_tbl <- flip_tbl tmp_store.tbl + | _ -> + let process_method () = let i = replace_method f.p_fn_name f' in + Hashtbl.add name_tbls.methods_tmps_tbl i (flip_tbl tmp_store.tbl) in + match f.p_fn_visibility with + | None -> process_method () + (* default to public *) + | Some v -> match v with + | Private -> let i = replace_func f.p_fn_name f' in + Hashtbl.add name_tbls.funcs_tmps_tbl i (flip_tbl tmp_store.tbl) + | Public -> process_method () + +(* We're not really using these, but leaving them here in case we ever want to + * inculde struct typedefs as a language feature of minic *) +let process_type t = let _ = match t with + | Pstruct _ | Punion _ -> match_type t + | _ -> raise_msg "invalid type declaration" in () + +let add_name_tables () = + name_tbls.funcs_tbl <- flip_tbl func_store.tbl; + name_tbls.methods_tbl <- flip_tbl method_store.tbl; + name_tbls.vars_tbl <- flip_tbl glob_store.tbl + +let map_store store = + let f tbl name = + let (i, t) = Hashtbl.find tbl name in + D.Coq_pair (i, t) in + coqlist_of_list (List.map (f store.tbl) store.lst) + +let typecheck parsed = + let preprocess_declaration = function + | Ptype_decl t -> () + | Pvar_decl v -> () + | Pfunc f -> preprocess_func f in + let process_declaration = function + | Ptype_decl t -> process_type t + | Pvar_decl v -> process_glob glob_store v + | Pfunc f -> process_func f in + let () = List.iter preprocess_declaration parsed in + let () = List.iter process_declaration parsed in + add_name_tables (); + let vars = map_store glob_store in + let funcs = map_store func_store in + let methods = map_store method_store in + let constructor = coq_option !constructor in + name_tbls, Genv.new_genv vars funcs methods constructor diff --git a/unittests/.gitignore b/unittests/.gitignore new file mode 100644 index 0000000..f8d87e3 --- /dev/null +++ b/unittests/.gitignore @@ -0,0 +1,6 @@ +amm_with_addresses.ds +erc20Token0.ds +erc20Token1.ds +liquidityToken.ds +node_modules +build diff --git a/unittests/README.md b/unittests/README.md index 4b2b48a..93340bd 100644 --- a/unittests/README.md +++ b/unittests/README.md @@ -2,4 +2,5 @@ To run the tests: 1. `npm install` to get Node.js dependencies 2. `./start_conflux.sh` in a separate terminal, and start Ganache (for Ethereum tests) - Be sure to clone the `conflux-rust` repo, and within `start_conflux.sh`, set `conflux` to the location of your repo clone. + - Ganache can be run with `npm run ganache` if `npm install` was used to install dependencies. 3. `./run_unittests.sh` diff --git a/unittests/amm_test.ts b/unittests/amm_test.ts new file mode 100644 index 0000000..d909d0f --- /dev/null +++ b/unittests/amm_test.ts @@ -0,0 +1,85 @@ +import { Chain } from './chain' +import { ethers } from 'ethers' +import { execSync } from 'child_process'; + +export async function runTest(chains: Chain[]) { + const [ + liquidity_chain, + erc0_chain, + erc1_chain, + amm_chain + ] = chains + + await Promise.all([ + erc0_chain.deployContract('./erc20Token0'), + erc1_chain.deployContract('./erc20Token1'), + liquidity_chain.deployContract('./liquidityToken'), + ]) + + const addresses = [ + erc0_chain.getContractAddress(), + erc1_chain.getContractAddress(), + liquidity_chain.getContractAddress(), + ] + + console.log(addresses.join(" ")) + execSync('echo ' + addresses.join(" ") + ' | ./replace_addresses.sh') + await amm_chain.deployContract('./amm_with_addresses') + + const amm_address = amm_chain.getContractAddress() + const liquidity_provider = ethers.Wallet.createRandom() + const user = ethers.Wallet.createRandom() + + try { + let response = await erc0_chain.callMethod("balanceOf", [await erc0_chain.getAccountAddress()]) + console.log("response (balanceOf):", response.toString()) + + response = await erc0_chain.callMethod("transfer", [liquidity_provider.address, 60000]) + console.log("response (transfer erc0 -> liq_prov):", response) + + response = await erc0_chain.callMethod("balanceOf", [liquidity_provider.address]) + console.log("response (balanceOf liq_prov in erc0):", response.toString()) + + response = await erc1_chain.callMethod("transfer", [liquidity_provider.address, 40000]) + console.log("response (transfer erc1 -> liq_prov):", response) + + response = await erc0_chain.callMethod("transfer", [amm_address, 500]) + console.log("response (transfer erc0 -> AMM):", response) + + response = await erc1_chain.callMethod("transfer", [amm_address, 500]) + console.log("response (transfer erc1 -> AMM):", response) + + response = await amm_chain.callMethod("mint", [liquidity_provider.address]) + console.log("response (mint):", response) + + // Pretend the user has some erc0 that they want to transfer to AMM + // (actually going straight from erc0 contract to AMM) + response = await erc0_chain.callMethod("transfer", [amm_address, 200]) + console.log("response (transfer erc0 -> AMM):", response) + + response = await amm_chain.callMethod("getReserve0") + console.log("response (getReserve0):", response.toString()) + + response = await amm_chain.callMethod("getBalance0") + console.log("response (getBalance0):", response.toString()) + + response = await erc0_chain.callMethod("balanceOf", [user.address]) + console.log("response (balanceOf user in erc0):", response.toString()) + response = await erc1_chain.callMethod("balanceOf", [user.address]) + console.log("response (balanceOf user in erc1):", response.toString()) + + response = await amm_chain.callMethod("simpleSwap0", [user.address]) + console.log("response (simpleSwap0):", response) + + response = await erc0_chain.callMethod("balanceOf", [user.address]) + console.log("response (balanceOf user in erc0):", response.toString()) + response = await erc1_chain.callMethod("balanceOf", [user.address]) + console.log("response (balanceOf user in erc1):", response.toString()) + + } + catch (e) { + console.log(e) + } +} + +export default { runTest } diff --git a/unittests/builtins_test.ts b/unittests/builtins_test.ts index 983a5ca..ddeb09c 100644 --- a/unittests/builtins_test.ts +++ b/unittests/builtins_test.ts @@ -31,8 +31,8 @@ export async function runTest(chain: Chain) { ]) printTest("address", _address == chain.getContractAddress()); - printTest("origin", _origin == chain.getAccountAddress()); - printTest("caller", _caller == chain.getAccountAddress()); + printTest("origin", _origin == await chain.getAccountAddress()); + printTest("caller", _caller == await chain.getAccountAddress()); printTest("callvalue", _callvalue == _value); printTest("timestamp", _timestamp == await chain.getBlockTimestamp(blockNumber)); printTest("number", _number == blockNumber); diff --git a/unittests/chain.ts b/unittests/chain.ts index 8f1c58d..07639a9 100644 --- a/unittests/chain.ts +++ b/unittests/chain.ts @@ -1,9 +1,9 @@ export interface Chain { - deployContract(jsonFilename: string, constructorArgs?: any[]): void + deployContract(jsonFilename: string, constructorArgs?: any[]): Promise callMethod(func: string, args?: any[], options?: {}): any getContractAddress(): string - getAccountAddress(): string + getAccountAddress(): Promise getBlockTimestamp(blockNumber: number): Promise getBlockParentHash(blockNumber: number): Promise getEvents(): Promise diff --git a/unittests/clean-json.py b/unittests/clean-json.py deleted file mode 100644 index 07c8e32..0000000 --- a/unittests/clean-json.py +++ /dev/null @@ -1,11 +0,0 @@ -import sys -import json - -s = sys.stdin.read() -obj = json.loads(s) -contracts = obj[list(obj)[0]] -first_contract = contracts[list(contracts)[0]] -first_contract['bytecode'] = '0x' + first_contract['bin'] -del first_contract['bin'] -first_contract['abi'] = json.loads(first_contract['abi']) -print(json.dumps(first_contract)) diff --git a/unittests/conflux-chain.ts b/unittests/conflux-chain.ts index 5ac9623..23d1b86 100644 --- a/unittests/conflux-chain.ts +++ b/unittests/conflux-chain.ts @@ -1,12 +1,12 @@ import { Conflux, Account } from 'js-conflux-sdk' import { Chain, Event } from './chain' -import { getJsonFilename } from './utils' +import { cleanCombined, getJsonFilename } from './utils' import _ from 'lodash' export class ConfluxChain implements Chain { private _conflux: Conflux - private _contract + private _contract: any private _genesisAccount: Account private _abi: any[] @@ -36,8 +36,9 @@ export class ConfluxChain implements Chain { return logger } - async deployContract(jsonFilename: string, constructorArgs=[]){ - const {abi, bytecode} = require(getJsonFilename(jsonFilename)) + async deployContract(jsonFilename: string, constructorArgs=[]) { + const combined = require(getJsonFilename(jsonFilename)) + const {abi, bytecode} = cleanCombined(combined) this._contract = this._conflux.Contract({ abi, bytecode }) this._abi = abi @@ -81,13 +82,16 @@ export class ConfluxChain implements Chain { getContractAddress(): string { return this._contract.address } - getAccountAddress(): string { + + async getAccountAddress(): Promise { return this._genesisAccount.toString() } + async getBlockTimestamp(blockNumber: number): Promise { let receipt: any = await this._conflux.getBlockByEpochNumber(blockNumber) return receipt.timestamp } + async getBlockParentHash(blockNumber: number): Promise { let receipt: any = await this._conflux.getBlockByEpochNumber(blockNumber) return receipt.parentHash @@ -116,6 +120,6 @@ export class ConfluxChain implements Chain { } interface Logger { - info() - error() + info(): any + error(): any } diff --git a/unittests/contracts/address.ds b/unittests/contracts/address.ds new file mode 100644 index 0000000..035f501 --- /dev/null +++ b/unittests/contracts/address.ds @@ -0,0 +1,28 @@ +object signature SimpleAddressInterface = { + init : unit -> unit; + const get : unit -> address; + set : address -> unit +} + +object SimpleAddress : SimpleAddressInterface { + let _val : address := address(0xefca) + + let init () = + _val := address(0xefca); + () + + let get () = + let val = _val in + val + + let set x = + _val := x +} + +layer signature SIMPLEADDRESSCONTRACTSig = { + simpleAddress : SimpleAddressInterface +} + +layer SIMPLEADDRESSCONTRACT : [{}] SIMPLEADDRESSCONTRACTSig = { + simpleAddress = SimpleAddress +} diff --git a/unittests/contracts/amm.ds b/unittests/contracts/amm.ds new file mode 100644 index 0000000..b5bb763 --- /dev/null +++ b/unittests/contracts/amm.ds @@ -0,0 +1,401 @@ +(* TODO: add emit Events *) + +(* PROOFS: + 1. change re-defined layers + 2. balance tracks reserve + 3. increasing k -> algebra + 4. no path independence, splitting trades is strictly more expensive -> algebra, appendix D + 5. liquidity token economics is right, i.e., UNI-V2 tracks the actual share of pool --> actually, where is the trading of UNI-V2 Pair tokens? What is its value? + 6. MOD: slippage control + 7. ... cost of manipulation (hard) +*) + +object signature ERC20Interface = { + constructor : unit -> unit; + mint : address * int -> unit; + burn : address * int -> unit; + totalSupply : unit -> int; + balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object LiquidityToken : ERC20Interface { + let _totalSupply : int := 0 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + balances[msg_sender] := 100000 + + let mint (toA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply + value; + let to_bal = balances[toA] in + balances[toA] := to_bal + value + + let burn (fromA, value) = + let totalSupply = _totalSupply in + _totalSupply := totalSupply - value; + let from_bal = balances[fromA] in + balances[fromA] := from_bal - value + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +object signature FixedERC20Interface = { + constructor : unit -> unit; + const totalSupply : unit -> int; + const balanceOf : address -> int; + transfer : address * int -> bool; + approve : address * int -> bool; + transferFrom : address * address * int -> bool +} + +object FixedSupplyToken : FixedERC20Interface { + let _totalSupply : int := 100000 + let balances : mapping[address] int := mapping_init + let allowances : mapping[address] mapping[address] int := mapping_init + + let constructor () = + _totalSupply := 100000; + balances[msg_sender] := 100000 + + let totalSupply () = + let bal0 = balances[address(0x0)] in + let totalSupply = _totalSupply in + let resultU = totalSupply - bal0 in + resultU + + let balanceOf tokenOwner = + let bal = balances[tokenOwner] in + let resultU = bal in + resultU + + let transfer(toA, tokens) = + let fromA = msg_sender in + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + assert (fromA <> toA /\ from_bal >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU + + let approve (spender, tokens) = + allowances[msg_sender][spender] := tokens; + let resultU = true in + resultU + + let transferFrom (fromA, toA, tokens) = + let from_bal = balances[fromA] in + let to_bal = balances[toA] in + let allowed = allowances[fromA][toA] in + assert (fromA <> toA /\ from_bal >= tokens /\ allowed >= tokens); + balances[fromA] := from_bal-tokens; + balances[toA] := to_bal+tokens; + let resultU = true in + resultU +} + +(* this AMM supports ERC20 tokens, ETH can be incorporated as WETH for trading *) +object signature AMMInterface = { + constructor : unit -> unit; + (* swap : int * int * address -> unit; *) + + coarsed simpleSwap0 : address -> int; (* transfer token0 to amm, results number of token1 swapped *) + + (* router *) + (* swapTokensForExactTokens + swapExactTokensForTokens + addLiquidity (* why restrain the liquidity to not depreciating an asset too much? *) + removeLiquidity *) + + const getReserve0 : unit -> int; + const getReserve1 : unit -> int; + const getBalance0 : unit -> int; + + (* fund management functions *) + coarsed mint : address -> unit; (* provide liquidity to this pair *) + coarsed burn : address -> unit; (* remove liquidity from this pair *) + + coarsed skim : address -> unit; (* arbitrage to force balances to match reserves *) + coarsed sync : unit -> unit; (* force reserves to match balances *) + k : unit -> int; (* get the constant product *) + + (* oracles *) + (* we do not provide Uniswap V2 version of smoothed oracles at this moment *) + quote0 : int -> int; (* given some amount of an asset and pair reserves, returns an equivalent amount of the other asset *) + getAmountOut0 : int -> int; (* given an input amount of an asset and pair reserves, returns the maximum output amount of the other asset *) + getAmountIn0 : int -> int; (* given an output amount of an asset and pair reserves, returns a required input amount of the other asset *) +} + +(* assuming 0.3% txn fee, excluding the 0.05% additional fee in Uniswap V2 design *) +(* we do not support liquidity tokens at the moment *) +object AutomatedMarketMaker + ( liquidityToken: ERC20Interface, + erc20Token0: FixedERC20Interface, + erc20Token1: FixedERC20Interface) : AMMInterface { + + let _token0 : address := address(0xdac17f958d2ee523a2206206994597c13d831ec7) (* USDT *) + let _token1 : address := address(0x6b175474e89094c44da98b954eedeac495271d0f) (* DAI *) + let _owner : address := address(0x0) + let _reserve0 : int := 0 + let _reserve1 : int := 0 + let blockTimestampLast : int := 0 (* new oracle *) + let price0CumulativeLast : int := 0 (* new oracle *) + let price1CumulativeLast : int := 0 (* new oracle *) + let kLast : int := 0 (* reserve0 * reserve1, as of immediately after the most recent liquidity event *) + + let constructor () = + _owner := msg_sender + + let getReserve0 () = + let reserve0 = _reserve0 in + reserve0 + + let getReserve1 () = + let reserve1 = _reserve1 in + reserve1 + + let getBalance0 () = + let balance0 = erc20Token0.balanceOf(this_address) in + balance0 + + (* transfer token0 and token1 to this contract, then mint liquidity tokens *) + let mint (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0 = balance0 - reserve0 in + let amount1 = balance1 - reserve1 in + (* update reserve 0 and reserve 1 *) + let totalSupply = liquidityToken.totalSupply() in + let liquidity = if totalSupply = 0 + then + begin + liquidityToken.mint(address(0x0), 1000); (* lock the first liquidity tokens*) + amount0 * amount1 - 1000 (* in Uniswap this should be square-rooted *) + end + else + let x = amount0 * totalSupply / reserve0 in + let y = amount1 * totalSupply / reserve1 in + if x < y then x else y + in + assert (liquidity > 0); + liquidityToken.mint(toA, liquidity); + _reserve0 := balance0; + _reserve1 := balance1 + + (* need toA first transfer liquidity tokens to this contract, then operate *) + (* it does not use ERC20 style allowances, just direct transfer *) + let burn (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let liquidity = liquidityToken.balanceOf(this_address) in + let totalSupply = liquidityToken.totalSupply() in + let amount0 = liquidity * balance0 / totalSupply in + let amount1 = liquidity * balance1 / totalSupply in + assert (amount0 > 0 /\ amount1 > 0); + liquidityToken.burn(this_address, liquidity); + let success = erc20Token0.transfer(toA, amount0) in + assert (success); + let success = erc20Token1.transfer(toA, amount1) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + (* adhere to the stanford paper formalization, targeted for vanilla proof *) + let simpleSwap0 (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = balance0 - reserve0 in + let token0 = _token0 in + let token1 = _token1 in + assert (toA <> token0 /\ toA <> token1); + assert (amount0In > 0); + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let result = numerator / denominator in + let success = erc20Token1.transfer(toA, result) in + assert (success); + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1; + let resultU = result in + resultU + + (* derives from Uniswap V2 flashswap, but not flashswap since DeepSEA does not allow control-flow aggregation *) + (* first transfer tokens to this contract, then call swap to swap them *) + (* notice this does not refund, so possibly not optimal rate *) + (* let swap (amount0Out, amount1Out, toA) = + assert (amount0Out > 0 \/ amount1Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let token0 = _token0 in + let token1 = _token1 in + assert (amount0Out < reserve0 /\ amount1Out < reserve1); + assert (toA <> token0 /\ toA <> token1); + (* check the assets satisfy what is wanted *) + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + (* we can prove that balance > reserve here *) + let amount0In = helper.getAmountIn(balance0, 0, reserve0) in (* balance0 - reserve0 *) + let amount1In = helper.getAmountIn(balance1, 0, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0 - amount0Out, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1 - amount1Out, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); + + (* sending money *) + let _ = if amount0Out > 0 then + let success = erc20Token0.transfer(toA, amount0Out) in + assert (success); + () + else + () + in + let _ = if amount1Out > 0 then + let success = erc20Token1.transfer(toA, amount1Out) in + assert (success); + () + else + () + in + (* flash swap style: sent money (amount0Out, amount1Out) to msg_sender, now wait to get back *) + (* after got money back *) + (* let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let amount0In = helper.getAmountIn(balance0, amount0Out, reserve0) in + let amount1In = helper.getAmountIn(balance1, amount1Out, reserve1) in + assert (amount0In > 0 \/ amount1In > 0); + let balance0Adjusted = helper.getBalanceAdjusted(balance0, amount0In) in + let balance1Adjusted = helper.getBalanceAdjusted(balance1, amount1In) in + assert (balance0Adjusted * balance1Adjusted > reserve0 * reserve1 * 1000 * 1000); *) + (* update the reserves, also update the oracle if possible *) + _reserve0 := balance0; + _reserve1 := balance1 *) + + let skim (toA) = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + let skim0 = balance0 - reserve0 in + let skim1 = balance1 - reserve1 in + let success = erc20Token0.transfer(toA, skim0) in + assert (success); + let success = erc20Token1.transfer(toA, skim1) in + assert (success) + + let sync () = + let balance0 = erc20Token0.balanceOf(this_address) in + let balance1 = erc20Token1.balanceOf(this_address) in + _reserve0 := balance0; + _reserve1 := balance1 + + let k () = + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + let resultU = reserve0 * reserve1 in + resultU + + (* Uniswap did not use Q112 library to do this computation *) + + (* given amount0 of token0, the equivalent value in token1 *) + let quote0 (amount0) = + assert (amount0 > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); (* prove that this can be dropped? *) + let resultU = amount0 * reserve1 / reserve0 in + resultU + + let getAmountOut0 (amount0In) = + assert (amount0In > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let amountInWithFee = amount0In * 997 in + let numerator = amountInWithFee * reserve1 in + let denominator = reserve0 * 1000 + amountInWithFee in + let resultU = numerator / denominator in + resultU + + let getAmountIn0 (amount0Out) = + assert (amount0Out > 0); + let reserve0 = _reserve0 in + let reserve1 = _reserve1 in + assert (reserve0 > 0 /\ reserve1 > 0); + let numerator = reserve1 * amount0Out * 1000 in + let denominator = (reserve0 - amount0Out) * 997 in + let resultU = (numerator / denominator) + 1 in + resultU +} + +layer signature AMMLibSig = { + erc20Token0 : FixedERC20Interface; + erc20Token1 : FixedERC20Interface; + liquidityToken : ERC20Interface; +} + +layer AMMLIB : [{}] AMMLibSig = { + erc20Token0 = address(0x10033) <: FixedSupplyToken; + erc20Token1 = address(0x10032) <: (clone FixedSupplyToken); (* cloned an object *) + liquidityToken = address(0x10031) <: LiquidityToken; +} + +layer signature AMMSig = { + amm : AMMInterface +} + +layer AMM : [AMMLibSig] AMMSig = { + amm = address(0x10030) <: AutomatedMarketMaker +} + +layer COMPLETE = AMM @ AMMLIB diff --git a/unittests/contracts/asserts.ds b/unittests/contracts/asserts.ds new file mode 100755 index 0000000..4b72626 --- /dev/null +++ b/unittests/contracts/asserts.ds @@ -0,0 +1,34 @@ +object signature OS = { + f : unit -> unit; + g : unit -> unit; + h : unit -> unit; + i : unit -> unit; + j : unit -> unit +} + +object O : OS { + (* should succeed *) + let f () = + assert (2+2 = 4) + + (* should revert *) + let g () = + assert (2+2 = 5) + + (* should revert *) + let h () = + deny (2+2 = 4) + + (* should succeed *) + let i () = + deny (2+2 = 5) + + (* should revert *) + let j () = + fail + +} + +layer L = { + o = O +} diff --git a/unittests/contracts/builtins.ds b/unittests/contracts/builtins.ds new file mode 100755 index 0000000..eba79ce --- /dev/null +++ b/unittests/contracts/builtins.ds @@ -0,0 +1,60 @@ +object signature OS = { + f : unit -> bool; + + const get_address : unit -> address; + const get_origin : unit -> address; + const get_caller : unit -> address; + const get_callvalue : unit -> uint; + const get_coinbase : unit -> uint; + const get_timestamp : unit -> uint; + const get_number : unit -> uint; + (* const get_chainid : unit -> uint; *) (*Not supported by our ethereum test yet*) + (* const get_selfbalance : unit -> uint; *) + const get_balance_this : unit -> uint; + const get_blockhash_prev : unit -> uint +} + +object O : OS { + let _address : address := address(0) + let _origin : address := address(0) + let _caller : address := address(0) + let _callvalue : uint := 0u0 + let _coinbase : uint := 0u0 + let _timestamp : uint := 0u0 + let _number : uint := 0u0 + (* let _chainid : uint := 0u0 *) + (* let _selfbalance : uint := 0u0 *) + let _balance_this : uint := 0u0 + let _blockhash_prev : uint := 0u0 + + + let get_address () = let x = _address in x + let get_origin () = let x = _origin in x + let get_caller () = let x = _caller in x + let get_callvalue () = let x = _callvalue in x + let get_coinbase () = let x = _coinbase in x + let get_timestamp () = let x = _timestamp in x + let get_number () = let x = _number in x + (* let get_chainid () = let x = _chainid in x *) + (* let get_selfbalance () = let x = _selfbalance in x *) + let get_balance_this () = let x = _balance_this in x + let get_blockhash_prev () = let x = _blockhash_prev in x + + let f () = + _address := this_address; + _origin := tx_origin; + _caller := msg_sender; + _callvalue := msg_value; + _coinbase := block_coinbase; + _timestamp := block_timestamp; + _number := block_number; + (* _chainid := chain_id; *) + (* _selfbalance := self_balance; *) + _balance_this := balance(this_address); + _blockhash_prev := blockhash(block_number-0u1); + true + } + +layer L = { + o = O +} diff --git a/unittests/contracts/constructor.ds b/unittests/contracts/constructor.ds new file mode 100644 index 0000000..07f7247 --- /dev/null +++ b/unittests/contracts/constructor.ds @@ -0,0 +1,119 @@ +type addr := int + +type Test_struct = { + ts_int_val : int; + ts_addr_val : addr +} + +object signature LowerInterface = { + constructor : int -> unit; + test_val : unit -> unit; + const get_x : unit -> int; + const get_y : unit -> uint; + const get_temp_val : unit -> int +} + +object Lower : LowerInterface { + let x : int := 6 + let y : uint := 0u123 + let temp_val : int := 1 + + let constructor y = + let temp_test_val = 2 in + temp_val := temp_test_val; + x := y + + let test_val () = + let a = x in + assert(a = 6) + + let get_x () = + let _x = x in + _x + + let get_y () = + let _y = y in + _y + + let get_temp_val () = + let _temp_val = temp_val in + _temp_val +} + +layer LOWER : [ { } ] { lower : LowerInterface } = { + lower = Lower +} + +object signature UpperInterface = { + constructor : int -> unit; + const get : unit -> unit +} + +object Upper (lower : LowerInterface) : UpperInterface { + let int_val : int := 10 + let uint_val : uint := 0u10 + let bool_val : bool := false + let address_val : address := address(0xefca) + let mapping_val : mapping[address] uint := mapping_init + let array_val : array[64] int := array_init + let struct_val : Test_struct := { ts_int_val=5; ts_addr_val=6 } + let hash_val : hashvalue := null_hash + let constructor_int_val : int := 123 + let constructor_temp_intval : int := 1234 + + let constructor z = + let temp_test_int = 20 in + constructor_temp_intval := temp_test_int; + constructor_int_val := z + + let get () = + let a = uint_val in + assert(a = 0u10); + + let b = int_val in + assert(b = 10); + + let c = bool_val in + assert(c = false); + + let d = address_val in + assert(d = address(0xefca)); + + let e = mapping_val[msg_sender] in + assert(e = 0u0); + + let f = array_val[0] in + assert(f = 0); + + let i = struct_val.ts_int_val in + assert(i = 5); + + let j = struct_val.ts_addr_val in + assert(j = 6); + + let l = lower.get_y() in + assert(l = 0u123); + + (* in lower layer, the x value firstly initialize as 6, and update in constructor to 100 by input argument *) + let g = lower.get_x() in + assert(g = 100); + + (* the constructor code overwrite the initializer's value, should be 321 here by input argument *) + let k = constructor_int_val in + assert(k = 321); + + let o = constructor_temp_intval in + assert(o = 20); + + let m = lower.get_temp_val() in + assert(m = 2) + +} + +layer UPPER : [{ lower : LowerInterface }] + { upper : UpperInterface } += { + upper = Upper +} + +layer CONTRACT = UPPER @ LOWER diff --git a/unittests/contracts/constructor2.ds b/unittests/contracts/constructor2.ds new file mode 100644 index 0000000..e1f72df --- /dev/null +++ b/unittests/contracts/constructor2.ds @@ -0,0 +1,59 @@ +(* An example where the constructors take more than one argument. *) + +object signature LowerInterface = { + constructor : int * int -> unit; + test_lower : unit -> unit; +} + +object Lower : LowerInterface { + let x : int := 0 + let y : int := 0 + + let constructor (_x, _y) = + x := _x; + y := _y + + let test_lower () = + let _x = x in + let _y = y in + assert(_x = 6 /\ _y = 7) + +} + +layer LOWER : [ { } ] { lower : LowerInterface } = { + lower = Lower +} + +object signature UpperInterface = { + constructor : int * int -> unit; + test_upper : unit -> unit; + const get_w : unit -> int +} + +object Upper (lower : LowerInterface) : UpperInterface { + let z : int := 0 + let w : int := 0 + + let constructor (_z, _w) = + z := _z; + w := _w + + let test_upper () = + lower.test_lower(); + let _z = z in + let _w = w in + assert (_z = 8 /\ _w = 9); + () + + let get_w () = + let _w = w in + w +} + +layer UPPER : [{ lower : LowerInterface }] + { upper : UpperInterface } += { + upper = Upper +} + +layer CONTRACT = UPPER @ LOWER diff --git a/unittests/contracts/events.ds b/unittests/contracts/events.ds new file mode 100644 index 0000000..b035bee --- /dev/null +++ b/unittests/contracts/events.ds @@ -0,0 +1,37 @@ +event + | E0 + +event + | E1 (v : uint) + +event + E2 (v1 : uint) (v2 : bool) + +event + | IE1 (v : uint indexed) + | IE2 (v1 : uint) (v2 : bool indexed) + +event + | Transfer (fromA : uint indexed) (toA: uint indexed) (value: int) + | Approval (owner: uint indexed) (spender: uint indexed) (value: int) + +object signature OS = { + f : unit -> bool +} + +object O : OS { + let f () = + emit E0(); + emit E1(0u42); + emit E2(0u1, false); + emit IE1(0u13); + emit IE2(0u2, false); + emit Transfer(0u1, 0u2, 42); + emit Approval(0u1, 0u2, 42); + true + } + +layer L = { + o = O +} + diff --git a/unittests/contracts/forloop.ds b/unittests/contracts/forloop.ds new file mode 100755 index 0000000..7c0b8c8 --- /dev/null +++ b/unittests/contracts/forloop.ds @@ -0,0 +1,24 @@ +object signature OS = { + multiply : int * int -> unit +} + +object O : OS { + + let _val : int := 0 + + (* a can be negative; + b must be positive *) + let multiply (a, b) = + for i = 0 to b do + begin + let val = _val in + _val := val + a + end; + let val = _val in assert(val = a*b); + _val := 0 + +} + +layer L = { + o = O +} diff --git a/unittests/contracts/keccak256.ds b/unittests/contracts/keccak256.ds new file mode 100644 index 0000000..bb6f79f --- /dev/null +++ b/unittests/contracts/keccak256.ds @@ -0,0 +1,47 @@ +object signature OS = { + const hash1 : uint -> hashvalue; + const hash2 : uint * uint -> hashvalue; + const hash2_1 : uint * hashvalue -> hashvalue; + const hash2_2 : hashvalue * uint -> hashvalue; + const hash2_3 : hashvalue * hashvalue -> hashvalue; + + const hash1a : address -> hashvalue; + const hash2a : address * address -> hashvalue; + const hash2_1a : address * hashvalue -> hashvalue; +} + +object O : OS { + + let blah : hashvalue := null_hash + + let hash1 x = + keccak256(x) + + let hash2 (x, y) = + keccak256 (x, y) + + let hash2_1 (x, y) = + keccak256 (x, y) + + let hash2_2 (x, y) = + keccak256 (x, y) + + let hash2_3 (x, y) = + keccak256 (x, y) + + let hash1a x = + keccak256(x) + + let hash2a (x, y) = + keccak256 (x, y) + + let hash2_1a (x, y) = + keccak256 (x, y) + + } + + + +layer L = { + o = O +} diff --git a/unittests/contracts/layers1.ds b/unittests/contracts/layers1.ds new file mode 100755 index 0000000..2b4d8ec --- /dev/null +++ b/unittests/contracts/layers1.ds @@ -0,0 +1,47 @@ +object signature LowerInterface = { + incr : unit -> unit; + const get : unit -> int +} + +object Lower : LowerInterface { + let n : int := 0 + + let incr () = + let _n = n in + n := _n + 1 + + let get () = + let _n = n in + _n + +} + +layer LOWER : [ { } ] { lower : LowerInterface } = { + lower = Lower +} + +object signature UpperInterface = { + f : unit -> int; + const get : unit -> int +} + +object Upper (lower : LowerInterface) : UpperInterface { + let f () = + lower.incr(); + lower.incr(); + let n = lower.get() in + assert (n = 2); + n + + let get() = + lower.get() + +} + +layer UPPER : [{ lower : LowerInterface }] + { upper : UpperInterface } += { + upper = Upper +} + +layer CONTRACT = UPPER @ LOWER \ No newline at end of file diff --git a/unittests/contracts/match-int.ds b/unittests/contracts/match-int.ds new file mode 100755 index 0000000..3071366 --- /dev/null +++ b/unittests/contracts/match-int.ds @@ -0,0 +1,30 @@ +type Status [[int]] = + | Ok [[= 0]] + | Triggered [[= 1]] + | Finalized [[= 2]] + +object signature OS = { + f : unit -> int +} + +object O : OS { + let f () = + let x = Ok in + let v = match x with + | Ok => + 40 + | Triggered => + 41 + | Finalized => + 42 + end + in + begin + assert (v = 40); + v + end + } + +layer L = { + o = O +} diff --git a/unittests/contracts/simple_call.ds b/unittests/contracts/simple_call.ds new file mode 100644 index 0000000..6b854f5 --- /dev/null +++ b/unittests/contracts/simple_call.ds @@ -0,0 +1,83 @@ +object signature FixedERC20Interface = { + const getNum : unit -> int; + doubleNum : unit -> unit; + const sum : int * int -> int; +} + +object FixedSupplyToken : FixedERC20Interface { + let myNum : int := 5 + + let getNum () = + let num = myNum in + num + + let doubleNum () = + let num = myNum in + myNum := num * 2; + () + + let sum (num1, num2) = + let ret = (2 * num1) + num2 in + ret +} + +object signature AMMInterface = { + call_double : unit -> unit; + const call_get : unit -> int; + const call_sum : int * int -> int; + const stress_call : int * int -> int; +} + +object AutomatedMarketMaker (erc20Token0: FixedERC20Interface) : AMMInterface { + let call_double () = + erc20Token0.doubleNum() + + let call_get () = + let ret = erc20Token0.getNum() in + ret + + let call_sum (num1, num2) = + let ret = erc20Token0.sum(num1, num2) in + ret + + let stress_call (num1, num2) = + let temp1 = 1 in + let temp2 = 2 in + let temp3 = 3 in + let temp4 = 4 in + let temp5 = 5 in + (*let temp6 = 6 in (* breaks down *) + let temp7 = 7 in + let temp8 = 8 in + let temp9 = 9 in + let temp10 = 10 in + let temp11 = 11 in + let temp12 = 12 in + let temp13 = 13 in + let temp14 = 14 in + let temp15 = 15 in + let temp16 = 16 in + let temp16 = 16 in*) + let ret = erc20Token0.sum(num1, num2) in + ret +} + +(* CUT *) + +layer signature AMMLibSig = { + erc20Token0 : FixedERC20Interface; +} + +layer AMMLIB : [{}] AMMLibSig = { + erc20Token0 = address(0x10033) <: FixedSupplyToken; +} + +layer signature AMMSig = { + amm : AMMInterface +} + +layer AMM : [AMMLibSig] AMMSig = { + amm = address(0x10030) <: AutomatedMarketMaker +} + +layer COMPLETE = AMM @ AMMLIB diff --git a/unittests/contracts/struct.ds b/unittests/contracts/struct.ds new file mode 100755 index 0000000..45dbc82 --- /dev/null +++ b/unittests/contracts/struct.ds @@ -0,0 +1,34 @@ +type addr := int + +object signature OS = { + initialize : int * int -> unit; + hasOwner : unit -> unit; + hasSupply : int -> unit; +} + +type Token = { + totalSupply : int; + owner : addr +} + +object O : OS { + let token : Token := { totalSupply=0; owner=0 } + + let initialize (_totalSupply, _owner) = + token.totalSupply := _totalSupply; + token.owner := _owner; + () + + let hasOwner () = + let _owner = token.owner in + assert (_owner <> 0) + + let hasSupply (amount) = + let _totalSupply = token.totalSupply in + assert (_totalSupply >= amount) + +} + +layer L = { + o = O +} diff --git a/unittests/cut_amm.sh b/unittests/cut_amm.sh new file mode 100755 index 0000000..95b5f1a --- /dev/null +++ b/unittests/cut_amm.sh @@ -0,0 +1,18 @@ +#!/bin/bash + +amm="contracts/amm.ds" + +fixed_supply_token=' +layer L = { + o = FixedSupplyToken +}' + +liquidity_token=' +layer L = { + o = LiquidityToken +}' + +amm_cut=$( gsed '/\(\* CUT \*\)/Q' "$amm") + +echo "$amm_cut" "$fixed_supply_token" | tee erc20Token0.ds erc20Token1.ds >/dev/null +echo "$amm_cut" "$liquidity_token" > liquidityToken.ds diff --git a/unittests/ethereum-chain.ts b/unittests/ethereum-chain.ts index 2d8b7f7..1e2d5a0 100644 --- a/unittests/ethereum-chain.ts +++ b/unittests/ethereum-chain.ts @@ -1,12 +1,13 @@ import { ethers } from 'ethers' import { Chain, Event } from './chain' -import { getJsonFilename } from './utils' +import { cleanCombined, getJsonFilename } from './utils' import _ from 'lodash' export class EthereumChain implements Chain { private _provider: ethers.providers.JsonRpcProvider private _contract: ethers.Contract + private _contracts: ethers.Contract[] private _abi: any[] private _signer: ethers.providers.JsonRpcSigner private _creator: string @@ -16,12 +17,12 @@ export class EthereumChain implements Chain { this._provider = new ethers.providers.JsonRpcProvider(endpoint); this._signer = this._provider.getSigner(0); + this._contracts = [] } async deployContract(jsonFilename: string, constructorArgs=[]) { - this._creator = await this._signer.getAddress(); - - const {abi, bytecode} = require(getJsonFilename(jsonFilename)) + const combined = require(getJsonFilename(jsonFilename)) + const {abi, bytecode} = cleanCombined(combined) this._abi = abi let factory = new ethers.ContractFactory(abi, bytecode, this._signer); @@ -29,7 +30,19 @@ export class EthereumChain implements Chain { await this._contract.deployed(); } - async callRead(func: string, args) { + async deployContracts(jsonFilenames: string[], constructorArgs: any[][]) { + for (const i in jsonFilenames) { + const combined = require(getJsonFilename(jsonFilenames[i])) + const {abi, bytecode} = cleanCombined(combined) + + let factory = new ethers.ContractFactory(abi, bytecode, this._signer); + let contract = await (factory.deploy.apply(factory, constructorArgs[i])) + this._contracts.push(contract); + await this._contracts[i].deployed(); // TODO: parallelize this loop + } + } + + async callRead(func: string, args=[]) { return await this._contract[func](...args) } @@ -46,18 +59,26 @@ export class EthereumChain implements Chain { async callMethod(func: string, args=[], options={}) { if (this.isRead(func)) { - if (!_.isEmpty(options)) - throw Error("Options specified but ignored.") - return this.callRead(func, args) // returns the read value - } else - return this.callWrite(func, args, options) // returns the block number + if (!_.isEmpty(options)) + throw Error("Options specified but ignored.") + return this.callRead(func, args) // returns the read value + } else + return this.callWrite(func, args, options) // returns the block number } getContractAddress(): string { return this._contract.address } - getAccountAddress(): string { + getContractAddresses(): string[] { + return this._contracts.map(c => c.address) + } + + async getAccountAddress(): Promise { + if (typeof this._creator === 'undefined' || this._creator === null) { + this._creator = await this._signer.getAddress(); + } + return this._creator } diff --git a/unittests/genesis_secrets.txt b/unittests/genesis_secrets.txt deleted file mode 100644 index 93f308c..0000000 --- a/unittests/genesis_secrets.txt +++ /dev/null @@ -1 +0,0 @@ -aaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaa diff --git a/unittests/package-lock.json b/unittests/package-lock.json index 82d67cb..48c6e83 100644 --- a/unittests/package-lock.json +++ b/unittests/package-lock.json @@ -388,6 +388,18 @@ "@ethersproject/strings": "^5.0.4" } }, + "@types/js-yaml": { + "version": "4.0.0", + "resolved": "https://registry.npmjs.org/@types/js-yaml/-/js-yaml-4.0.0.tgz", + "integrity": "sha512-4vlpCM5KPCL5CfGmTbpjwVKbISRYhduEJvvUWsH5EB7QInhEj94XPZ3ts/9FPiLZFqYO0xoW4ZL8z2AabTGgJA==", + "dev": true + }, + "@types/lodash": { + "version": "4.14.167", + "resolved": "https://registry.npmjs.org/@types/lodash/-/lodash-4.14.167.tgz", + "integrity": "sha512-w7tQPjARrvdeBkX/Rwg95S592JwxqOjmms3zWQ0XZgSyxSLdzWaYH3vErBhdVS/lRBX7F8aBYcYJYTr5TMGOzw==", + "dev": true + }, "@types/node": { "version": "14.14.12", "resolved": "https://registry.npmjs.org/@types/node/-/node-14.14.12.tgz", @@ -1636,9 +1648,9 @@ } }, "typescript": { - "version": "4.1.2", - "resolved": "https://registry.npmjs.org/typescript/-/typescript-4.1.2.tgz", - "integrity": "sha512-thGloWsGH3SOxv1SoY7QojKi0tc+8FnOmiarEGMbd/lar7QOEd3hvlx3Fp5y6FlDUGl9L+pd4n2e+oToGMmhRQ==" + "version": "4.1.3", + "resolved": "https://registry.npmjs.org/typescript/-/typescript-4.1.3.tgz", + "integrity": "sha512-B3ZIOf1IKeH2ixgHhj6la6xdwR9QrLC5d1VKeCSY4tvkqhF2eqd9O7txNlS0PO3GrBAFIdr3L1ndNwteUbZLYg==" }, "utf-8-validate": { "version": "5.0.3", diff --git a/unittests/package.json b/unittests/package.json index a3cd11b..ef32ea8 100644 --- a/unittests/package.json +++ b/unittests/package.json @@ -13,5 +13,9 @@ "scripts": { "ganache": "ganache-cli", "tsc": "tsc" + }, + "devDependencies": { + "@types/js-yaml": "^4.0.0", + "@types/lodash": "^4.14.167" } } diff --git a/unittests/reference/README.md b/unittests/reference/README.md new file mode 100644 index 0000000..d0d85bc --- /dev/null +++ b/unittests/reference/README.md @@ -0,0 +1,3 @@ +This directory has Solidity programs doing the "same thing" as some of the unit +tests, they were created to check that our behaviour is abi-compatible with +Solidity. diff --git a/unittests/reference/events.sol b/unittests/reference/events.sol new file mode 100644 index 0000000..7345791 --- /dev/null +++ b/unittests/reference/events.sol @@ -0,0 +1,21 @@ +pragma solidity ^0.5.12; + +contract O { + event E1 (uint v); + event E2 (uint v1, bool v2); + event IE1 (uint indexed v); + event IE2 (uint v1, bool indexed v2); + event Transfer (uint indexed fromA, uint indexed toA, uint value); + event Approval (uint indexed owner, uint indexed spender, uint value); + + + function f() public { + emit E1(42); + emit E2(1, false); + emit IE1(13); + emit IE2(2, false); + emit Transfer(1, 2, 42); + emit Approval(1, 2, 42); + } + +} diff --git a/unittests/reference/keccak256.sol b/unittests/reference/keccak256.sol new file mode 100644 index 0000000..2c1f634 --- /dev/null +++ b/unittests/reference/keccak256.sol @@ -0,0 +1,16 @@ +pragma solidity ^0.5.12; + +contract O { + function hash1(uint x ) public pure returns (uint) { + return uint256(keccak256(abi.encodePacked(x))); + } + + function hash2(uint x, uint y) public pure returns (uint) { + return uint256(keccak256(abi.encodePacked(x,y))); + } + + function hash1a(address x) public pure returns (uint) { + return uint256(keccak256(abi.encodePacked(uint(x)))); + } + +} diff --git a/unittests/replace_addresses.sh b/unittests/replace_addresses.sh new file mode 100755 index 0000000..4af64f0 --- /dev/null +++ b/unittests/replace_addresses.sh @@ -0,0 +1,21 @@ +#!/bin/bash + +set -x + +dsc='../binaries/MacOS/dsc' +amm='contracts/amm.ds' +amm_basename=$(basename $amm) +amm_with_addresses=${amm_basename%.ds}_with_addresses.ds + +addresses=$( "$amm_with_addresses" + +"$dsc" "$amm_with_addresses" combined-json > "build/${amm_with_addresses%.ds}.json" diff --git a/unittests/run_unittests.sh b/unittests/run_unittests.sh index c14044d..53cdaa1 100755 --- a/unittests/run_unittests.sh +++ b/unittests/run_unittests.sh @@ -1,11 +1,15 @@ #!/bin/bash -edsger=../binaries/MacOS/dsc +dsc=../binaries/MacOS/dsc if [[ ! -d build ]]; then mkdir build; fi -for deepsea in *.ds; do - $edsger $deepsea combined-json | python3 clean-json.py > build/${deepsea%.ds}.json +./cut_amm.sh + +for deepsea in ./contracts/*.ds; do + basename=${deepsea##*/} + echo Compiling $basename + $dsc $deepsea combined-json > build/${basename%.ds}.json done npm run tsc diff --git a/unittests/start_conflux.sh b/unittests/start_conflux.sh index ba4d016..30b7b4c 100755 --- a/unittests/start_conflux.sh +++ b/unittests/start_conflux.sh @@ -3,7 +3,11 @@ # Set this to the location of the conflux repo. conflux=~/Documents/conflux-rust -cp ./tethys.toml ./genesis_secrets.txt $conflux/run +cp ./tethys.toml $conflux/run +genPrivateKey=aaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaa +genesis_secrets=$conflux/run/genesis_secrets.txt +if [[ -f $genesis_secrets ]]; then rm $genesis_secrets; fi +echo $genPrivateKey > $genesis_secrets cd $conflux/run ./clear_state.sh $conflux/target/release/conflux --config tethys.toml diff --git a/unittests/tests.ts b/unittests/tests.ts index 4ff0204..af69984 100644 --- a/unittests/tests.ts +++ b/unittests/tests.ts @@ -7,9 +7,10 @@ import fs from 'fs' import builtin from './builtins_test' import keccak256 from './keccak256_test' import events from './events_test' +import amm from './amm_test' async function main() { - const tests = yaml.safeLoad(fs.readFileSync('./tests.yml')) + const tests = yaml.load(fs.readFileSync('./tests.yml', 'utf8')) const chains = [ new EthereumChain, @@ -33,6 +34,9 @@ async function main() { await specialTests[t].runTest(chain) } + if (! (tests instanceof Object)) { + return // type assertion so `tests` can be looped over + } // These are tests that can be described in the `tests.json` file. They // either check whether a method call passes/reverts or the result of a // method's return value against a constant expected value. @@ -41,6 +45,10 @@ async function main() { await runTest(chain, t, tests[t].calls, tests[t].constructorArgs); } } + + amm.runTest(new Array(4).fill(null).map(() => (new EthereumChain))) + // amm.runTest(new Array(4).fill(null).map(() => (new ConfluxChain))) + } main() diff --git a/unittests/utils.ts b/unittests/utils.ts index ac9e306..1433c3f 100644 --- a/unittests/utils.ts +++ b/unittests/utils.ts @@ -14,3 +14,18 @@ export function padName(name: string): string { export function printTest(name: string, success: boolean) { console.log(padName(name), (success? "pass✅" : "fail❌")); } + +export function cleanCombined(combined: {}): abiBytecode { + const contracts = combined["contracts"] + const name = Object.keys(contracts)[0] + const contract = contracts[name] + return { + abi: JSON.parse(contract["abi"]), + bytecode: '0x' + contract["bin"] + } +} + +interface abiBytecode { + abi + bytecode: string +}