hugr_core/
import.rs

1//! Importing HUGR graphs from their `hugr-model` representation.
2//!
3//! **Warning**: This module is still under development and is expected to change.
4//! It is included in the library to allow for early experimentation, and for
5//! the core and model to converge incrementally.
6use std::sync::Arc;
7
8use crate::{
9    Direction, Hugr, HugrView, Node, Port,
10    extension::{
11        ExtensionId, ExtensionRegistry, SignatureError, resolution::ExtensionResolutionError,
12    },
13    hugr::{HugrMut, NodeMetadata},
14    ops::{
15        AliasDecl, AliasDefn, CFG, Call, CallIndirect, Case, Conditional, Const, DFG,
16        DataflowBlock, ExitBlock, FuncDecl, FuncDefn, Input, LoadConstant, LoadFunction, OpType,
17        OpaqueOp, Output, Tag, TailLoop, Value,
18        constant::{CustomConst, CustomSerialized, OpaqueValue},
19    },
20    package::Package,
21    std_extensions::{
22        arithmetic::{float_types::ConstF64, int_types::ConstInt},
23        collections::array::ArrayValue,
24    },
25    types::{
26        CustomType, FuncTypeBase, MaybeRV, PolyFuncType, PolyFuncTypeBase, RowVariable, Signature,
27        Term, Type, TypeArg, TypeBase, TypeBound, TypeEnum, TypeName, TypeRow,
28        type_param::{SeqPart, TypeParam},
29        type_row::TypeRowBase,
30    },
31};
32use fxhash::FxHashMap;
33use hugr_model::v0 as model;
34use hugr_model::v0::table;
35use itertools::{Either, Itertools};
36use smol_str::{SmolStr, ToSmolStr};
37use thiserror::Error;
38
39/// An error that can occur during import.
40#[derive(Debug, Clone, Error)]
41#[error("failed to import hugr")]
42pub struct ImportError(#[from] ImportErrorInner);
43
44#[derive(Debug, Clone, Error)]
45enum ImportErrorInner {
46    /// The model contains a feature that is not supported by the importer yet.
47    /// Errors of this kind are expected to be removed as the model format and
48    /// the core HUGR representation converge.
49    #[error("currently unsupported: {0}")]
50    Unsupported(String),
51
52    /// The model contains implicit information that has not yet been inferred.
53    /// This includes wildcards and application of functions with implicit parameters.
54    #[error("uninferred implicit: {0}")]
55    Uninferred(String),
56
57    /// The model is not well-formed.
58    #[error("{0}")]
59    Invalid(String),
60
61    /// An error with additional context.
62    #[error("import failed in context: {1}")]
63    Context(#[source] Box<ImportErrorInner>, String),
64
65    /// A signature mismatch was detected during import.
66    #[error("signature error")]
67    Signature(#[from] SignatureError),
68
69    /// An error relating to the loaded extension registry.
70    #[error("extension error")]
71    Extension(#[from] ExtensionError),
72
73    /// Incorrect order hints.
74    #[error("incorrect order hint")]
75    OrderHint(#[from] OrderHintError),
76
77    /// Extension resolution.
78    #[error("extension resolution error")]
79    ExtensionResolution(#[from] ExtensionResolutionError),
80}
81
82#[derive(Debug, Clone, Error)]
83enum ExtensionError {
84    /// An extension is missing.
85    #[error("Importing the hugr requires extension {missing_ext}, which was not found in the registry. The available extensions are: [{}]",
86            available.iter().map(std::string::ToString::to_string).collect::<Vec<_>>().join(", "))]
87    Missing {
88        /// The missing extension.
89        missing_ext: ExtensionId,
90        /// The available extensions in the registry.
91        available: Vec<ExtensionId>,
92    },
93
94    /// An extension type is missing.
95    #[error(
96        "Importing the hugr requires extension {ext} to have a type named {name}, but it was not found."
97    )]
98    MissingType {
99        /// The extension that is missing the type.
100        ext: ExtensionId,
101        /// The name of the missing type.
102        name: TypeName,
103    },
104}
105
106impl From<ExtensionError> for ImportError {
107    fn from(value: ExtensionError) -> Self {
108        Self::from(ImportErrorInner::from(value))
109    }
110}
111
112/// Import error caused by incorrect order hints.
113#[derive(Debug, Clone, Error)]
114enum OrderHintError {
115    /// Duplicate order hint key in the same region.
116    #[error("duplicate order hint key {0}")]
117    DuplicateKey(table::RegionId, u64),
118    /// Order hint including a key not defined in the region.
119    #[error("order hint with unknown key {0}")]
120    UnknownKey(u64),
121    /// Order hint involving a node with no order port.
122    #[error("order hint on node with no order port: {0}")]
123    NoOrderPort(table::NodeId),
124}
125
126impl From<OrderHintError> for ImportError {
127    fn from(value: OrderHintError) -> Self {
128        Self::from(ImportErrorInner::from(value))
129    }
130}
131
132/// Helper macro to create an `ImportError::Unsupported` error with a formatted message.
133macro_rules! error_unsupported {
134    ($($e:expr),*) => { ImportError(ImportErrorInner::Unsupported(format!($($e),*))) }
135}
136
137/// Helper macro to create an `ImportError::Uninferred` error with a formatted message.
138macro_rules! error_uninferred {
139    ($($e:expr),*) => { ImportError(ImportErrorInner::Uninferred(format!($($e),*))) }
140}
141
142/// Helper macro to create an `ImportError::Invalid` error with a formatted message.
143macro_rules! error_invalid {
144    ($($e:expr),*) => { ImportError(ImportErrorInner::Invalid(format!($($e),*))) }
145}
146
147/// Helper macro to create an `ImportError::Context` error with a formatted message.
148macro_rules! error_context {
149    ($err:expr, $($e:expr),*) => {
150        {
151            let ImportError(__err) = $err;
152            ImportError(ImportErrorInner::Context(Box::new(__err), format!($($e),*)))
153        }
154    }
155}
156
157/// Import a [`Package`] from its model representation.
158pub fn import_package(
159    package: &table::Package,
160    extensions: &ExtensionRegistry,
161) -> Result<Package, ImportError> {
162    let modules = package
163        .modules
164        .iter()
165        .map(|module| import_hugr(module, extensions))
166        .collect::<Result<Vec<_>, _>>()?;
167
168    // This does not panic since the import already requires a module root.
169    let package = Package::new(modules);
170    Ok(package)
171}
172
173/// Import a [`Hugr`] module from its model representation.
174pub fn import_hugr(
175    module: &table::Module,
176    extensions: &ExtensionRegistry,
177) -> Result<Hugr, ImportError> {
178    // TODO: Module should know about the number of edges, so that we can use a vector here.
179    // For now we use a hashmap, which will be slower.
180    let mut ctx = Context {
181        module,
182        hugr: Hugr::new(),
183        link_ports: FxHashMap::default(),
184        static_edges: Vec::new(),
185        extensions,
186        nodes: FxHashMap::default(),
187        local_vars: FxHashMap::default(),
188        custom_name_cache: FxHashMap::default(),
189        region_scope: table::RegionId::default(),
190    };
191
192    ctx.import_root()?;
193    ctx.link_ports()?;
194    ctx.link_static_ports()?;
195
196    ctx.hugr
197        .resolve_extension_defs(extensions)
198        .map_err(ImportErrorInner::ExtensionResolution)?;
199
200    Ok(ctx.hugr)
201}
202
203struct Context<'a> {
204    /// The module being imported.
205    module: &'a table::Module<'a>,
206
207    /// The HUGR graph being constructed.
208    hugr: Hugr,
209
210    /// The ports that are part of each link. This is used to connect the ports at the end of the
211    /// import process.
212    link_ports: FxHashMap<(table::RegionId, table::LinkIndex), Vec<(Node, Port)>>,
213
214    /// Pairs of nodes that should be connected by a static edge.
215    /// These are collected during the import process and connected at the end.
216    static_edges: Vec<(table::NodeId, table::NodeId)>,
217
218    /// The ambient extension registry to use for importing.
219    extensions: &'a ExtensionRegistry,
220
221    /// A map from `NodeId` to the imported `Node`.
222    nodes: FxHashMap<table::NodeId, Node>,
223
224    local_vars: FxHashMap<table::VarId, LocalVar>,
225
226    custom_name_cache: FxHashMap<&'a str, (ExtensionId, SmolStr)>,
227
228    region_scope: table::RegionId,
229}
230
231impl<'a> Context<'a> {
232    /// Get the signature of the node with the given `NodeId`.
233    fn get_node_signature(&mut self, node: table::NodeId) -> Result<Signature, ImportError> {
234        let node_data = self.get_node(node)?;
235        let signature = node_data
236            .signature
237            .ok_or_else(|| error_uninferred!("node signature"))?;
238        self.import_func_type(signature)
239    }
240
241    /// Get the node with the given `NodeId`, or return an error if it does not exist.
242    #[inline]
243    fn get_node(&self, node_id: table::NodeId) -> Result<&'a table::Node<'a>, ImportError> {
244        self.module
245            .get_node(node_id)
246            .ok_or_else(|| error_invalid!("unknown node {}", node_id))
247    }
248
249    /// Get the term with the given `TermId`, or return an error if it does not exist.
250    #[inline]
251    fn get_term(&self, term_id: table::TermId) -> Result<&'a table::Term<'a>, ImportError> {
252        self.module
253            .get_term(term_id)
254            .ok_or_else(|| error_invalid!("unknown term {}", term_id))
255    }
256
257    /// Get the region with the given `RegionId`, or return an error if it does not exist.
258    #[inline]
259    fn get_region(&self, region_id: table::RegionId) -> Result<&'a table::Region<'a>, ImportError> {
260        self.module
261            .get_region(region_id)
262            .ok_or_else(|| error_invalid!("unknown region {}", region_id))
263    }
264
265    fn make_node(
266        &mut self,
267        node_id: table::NodeId,
268        op: OpType,
269        parent: Node,
270    ) -> Result<Node, ImportError> {
271        let node = self.hugr.add_node_with_parent(parent, op);
272        self.nodes.insert(node_id, node);
273
274        let node_data = self.get_node(node_id)?;
275        self.record_links(node, Direction::Incoming, node_data.inputs);
276        self.record_links(node, Direction::Outgoing, node_data.outputs);
277
278        for meta_item in node_data.meta {
279            self.import_node_metadata(node, *meta_item)
280                .map_err(|err| error_context!(err, "node metadata"))?;
281        }
282
283        Ok(node)
284    }
285
286    fn import_node_metadata(
287        &mut self,
288        node: Node,
289        meta_item: table::TermId,
290    ) -> Result<(), ImportError> {
291        // Import the JSON metadata
292        if let Some([name_arg, json_arg]) = self.match_symbol(meta_item, model::COMPAT_META_JSON)? {
293            let table::Term::Literal(model::Literal::Str(name)) = self.get_term(name_arg)? else {
294                return Err(error_invalid!(
295                    "`{}` expects a string literal as its first argument",
296                    model::COMPAT_META_JSON
297                ));
298            };
299
300            let table::Term::Literal(model::Literal::Str(json_str)) = self.get_term(json_arg)?
301            else {
302                return Err(error_invalid!(
303                    "`{}` expects a string literal as its second argument",
304                    model::COMPAT_CONST_JSON
305                ));
306            };
307
308            let json_value: NodeMetadata = serde_json::from_str(json_str).map_err(|_| {
309                error_invalid!(
310                    "failed to parse JSON string for `{}` metadata",
311                    model::COMPAT_CONST_JSON
312                )
313            })?;
314
315            self.hugr.set_metadata(node, name, json_value);
316        }
317
318        // Set the entrypoint
319        if let Some([]) = self.match_symbol(meta_item, model::CORE_ENTRYPOINT)? {
320            self.hugr.set_entrypoint(node);
321        }
322
323        Ok(())
324    }
325
326    /// Associate links with the ports of the given node in the given direction.
327    fn record_links(&mut self, node: Node, direction: Direction, links: &'a [table::LinkIndex]) {
328        let optype = self.hugr.get_optype(node);
329        // NOTE: `OpType::port_count` copies the signature, which significantly slows down the import.
330        debug_assert!(links.len() <= optype.port_count(direction));
331
332        for (link, port) in links.iter().zip(self.hugr.node_ports(node, direction)) {
333            self.link_ports
334                .entry((self.region_scope, *link))
335                .or_default()
336                .push((node, port));
337        }
338    }
339
340    /// Link up the ports in the hugr graph, according to the connectivity information that
341    /// has been gathered in the `link_ports` map.
342    fn link_ports(&mut self) -> Result<(), ImportError> {
343        // For each edge, we group the ports by their direction. We reuse the `inputs` and
344        // `outputs` vectors to avoid unnecessary allocations.
345        let mut inputs = Vec::new();
346        let mut outputs = Vec::new();
347
348        for (link_id, link_ports) in std::mem::take(&mut self.link_ports) {
349            // Skip the edge if it doesn't have any ports.
350            if link_ports.is_empty() {
351                continue;
352            }
353
354            for (node, port) in link_ports {
355                match port.as_directed() {
356                    Either::Left(input) => inputs.push((node, input)),
357                    Either::Right(output) => outputs.push((node, output)),
358                }
359            }
360
361            match (inputs.as_slice(), outputs.as_slice()) {
362                ([], []) => {
363                    unreachable!();
364                }
365                (_, [output]) => {
366                    for (node, port) in &inputs {
367                        self.hugr.connect(output.0, output.1, *node, *port);
368                    }
369                }
370                ([input], _) => {
371                    for (node, port) in &outputs {
372                        self.hugr.connect(*node, *port, input.0, input.1);
373                    }
374                }
375                _ => {
376                    return Err(error_unsupported!(
377                        "link {:?} would require hyperedge",
378                        link_id
379                    ));
380                }
381            }
382
383            inputs.clear();
384            outputs.clear();
385        }
386
387        Ok(())
388    }
389
390    fn link_static_ports(&mut self) -> Result<(), ImportError> {
391        for (src_id, dst_id) in std::mem::take(&mut self.static_edges) {
392            // None of these lookups should fail given how we constructed `static_edges`.
393            let src = self.nodes[&src_id];
394            let dst = self.nodes[&dst_id];
395            let src_port = self.hugr.get_optype(src).static_output_port().unwrap();
396            let dst_port = self.hugr.get_optype(dst).static_input_port().unwrap();
397            self.hugr.connect(src, src_port, dst, dst_port);
398        }
399
400        Ok(())
401    }
402
403    fn get_symbol_name(&self, node_id: table::NodeId) -> Result<&'a str, ImportError> {
404        let node_data = self.get_node(node_id)?;
405        let name = node_data
406            .operation
407            .symbol()
408            .ok_or_else(|| error_invalid!("node {} is expected to be a symbol", node_id))?;
409        Ok(name)
410    }
411
412    fn get_func_signature(
413        &mut self,
414        func_node: table::NodeId,
415    ) -> Result<PolyFuncType, ImportError> {
416        let symbol = match self.get_node(func_node)?.operation {
417            table::Operation::DefineFunc(symbol) => symbol,
418            table::Operation::DeclareFunc(symbol) => symbol,
419            _ => {
420                return Err(error_invalid!(
421                    "node {} is expected to be a function declaration or definition",
422                    func_node
423                ));
424            }
425        };
426
427        self.import_poly_func_type(func_node, *symbol, |_, signature| Ok(signature))
428    }
429
430    /// Import the root region of the module.
431    fn import_root(&mut self) -> Result<(), ImportError> {
432        self.region_scope = self.module.root;
433        let region_data = self.get_region(self.module.root)?;
434
435        for node in region_data.children {
436            self.import_node(*node, self.hugr.entrypoint())?;
437        }
438
439        for meta_item in region_data.meta {
440            self.import_node_metadata(self.hugr.module_root(), *meta_item)?;
441        }
442
443        Ok(())
444    }
445
446    fn import_node(
447        &mut self,
448        node_id: table::NodeId,
449        parent: Node,
450    ) -> Result<Option<Node>, ImportError> {
451        let node_data = self.get_node(node_id)?;
452
453        let result = match node_data.operation {
454            table::Operation::Invalid => {
455                return Err(error_invalid!("tried to import an `invalid` operation"));
456            }
457
458            table::Operation::Dfg => Some(
459                self.import_node_dfg(node_id, parent, node_data)
460                    .map_err(|err| error_context!(err, "`dfg` node with id {}", node_id))?,
461            ),
462
463            table::Operation::Cfg => Some(
464                self.import_node_cfg(node_id, parent, node_data)
465                    .map_err(|err| error_context!(err, "`cfg` node with id {}", node_id))?,
466            ),
467
468            table::Operation::Block => Some(
469                self.import_node_block(node_id, parent)
470                    .map_err(|err| error_context!(err, "`block` node with id {}", node_id))?,
471            ),
472
473            table::Operation::DefineFunc(symbol) => Some(
474                self.import_node_define_func(node_id, symbol, node_data, parent)
475                    .map_err(|err| error_context!(err, "`define-func` node with id {}", node_id))?,
476            ),
477
478            table::Operation::DeclareFunc(symbol) => Some(
479                self.import_node_declare_func(node_id, symbol, parent)
480                    .map_err(|err| {
481                        error_context!(err, "`declare-func` node with id {}", node_id)
482                    })?,
483            ),
484
485            table::Operation::TailLoop => Some(
486                self.import_tail_loop(node_id, parent)
487                    .map_err(|err| error_context!(err, "`tail-loop` node with id {}", node_id))?,
488            ),
489
490            table::Operation::Conditional => Some(
491                self.import_conditional(node_id, parent)
492                    .map_err(|err| error_context!(err, "`cond` node with id {}", node_id))?,
493            ),
494
495            table::Operation::Custom(operation) => Some(
496                self.import_node_custom(node_id, operation, node_data, parent)
497                    .map_err(|err| error_context!(err, "custom node with id {}", node_id))?,
498            ),
499
500            table::Operation::DefineAlias(symbol, value) => Some(
501                self.import_node_define_alias(node_id, symbol, value, parent)
502                    .map_err(|err| {
503                        error_context!(err, "`define-alias` node with id {}", node_id)
504                    })?,
505            ),
506
507            table::Operation::DeclareAlias(symbol) => Some(
508                self.import_node_declare_alias(node_id, symbol, parent)
509                    .map_err(|err| {
510                        error_context!(err, "`declare-alias` node with id {}", node_id)
511                    })?,
512            ),
513
514            table::Operation::Import { .. } => None,
515
516            table::Operation::DeclareConstructor { .. } => None,
517            table::Operation::DeclareOperation { .. } => None,
518        };
519
520        Ok(result)
521    }
522
523    fn import_node_dfg(
524        &mut self,
525        node_id: table::NodeId,
526        parent: Node,
527        node_data: &'a table::Node<'a>,
528    ) -> Result<Node, ImportError> {
529        let signature = self
530            .get_node_signature(node_id)
531            .map_err(|err| error_context!(err, "node signature"))?;
532
533        let optype = OpType::DFG(DFG { signature });
534        let node = self.make_node(node_id, optype, parent)?;
535
536        let [region] = node_data.regions else {
537            return Err(error_invalid!("dfg region expects a single region"));
538        };
539
540        self.import_dfg_region(*region, node)?;
541        Ok(node)
542    }
543
544    fn import_node_cfg(
545        &mut self,
546        node_id: table::NodeId,
547        parent: Node,
548        node_data: &'a table::Node<'a>,
549    ) -> Result<Node, ImportError> {
550        let signature = self
551            .get_node_signature(node_id)
552            .map_err(|err| error_context!(err, "node signature"))?;
553
554        let optype = OpType::CFG(CFG { signature });
555        let node = self.make_node(node_id, optype, parent)?;
556
557        let [region] = node_data.regions else {
558            return Err(error_invalid!("cfg nodes expect a single region"));
559        };
560
561        self.import_cfg_region(*region, node)?;
562        Ok(node)
563    }
564
565    fn import_dfg_region(
566        &mut self,
567        region: table::RegionId,
568        node: Node,
569    ) -> Result<(), ImportError> {
570        let region_data = self.get_region(region)?;
571
572        let prev_region = self.region_scope;
573        if region_data.scope.is_some() {
574            self.region_scope = region;
575        }
576
577        if region_data.kind != model::RegionKind::DataFlow {
578            return Err(error_invalid!("expected dfg region"));
579        }
580
581        let signature = self
582            .import_func_type(
583                region_data
584                    .signature
585                    .ok_or_else(|| error_uninferred!("region signature"))?,
586            )
587            .map_err(|err| error_context!(err, "signature of dfg region with id {}", region))?;
588
589        // Create the input and output nodes
590        let input = self.hugr.add_node_with_parent(
591            node,
592            OpType::Input(Input {
593                types: signature.input,
594            }),
595        );
596        let output = self.hugr.add_node_with_parent(
597            node,
598            OpType::Output(Output {
599                types: signature.output,
600            }),
601        );
602
603        // Make sure that the ports of the input/output nodes are connected correctly
604        self.record_links(input, Direction::Outgoing, region_data.sources);
605        self.record_links(output, Direction::Incoming, region_data.targets);
606
607        for child in region_data.children {
608            self.import_node(*child, node)?;
609        }
610
611        self.create_order_edges(region, input, output)?;
612
613        for meta_item in region_data.meta {
614            self.import_node_metadata(node, *meta_item)?;
615        }
616
617        self.region_scope = prev_region;
618
619        Ok(())
620    }
621
622    /// Create order edges between nodes of a dataflow region based on order hint metadata.
623    ///
624    /// This method assumes that the nodes for the children of the region have already been imported.
625    fn create_order_edges(
626        &mut self,
627        region_id: table::RegionId,
628        input: Node,
629        output: Node,
630    ) -> Result<(), ImportError> {
631        let region_data = self.get_region(region_id)?;
632        debug_assert_eq!(region_data.kind, model::RegionKind::DataFlow);
633
634        // Collect order hint keys
635        // PERFORMANCE: It might be worthwhile to reuse the map to avoid allocations.
636        let mut order_keys = FxHashMap::<u64, Node>::default();
637
638        for child_id in region_data.children {
639            let child_data = self.get_node(*child_id)?;
640
641            for meta_id in child_data.meta {
642                let Some([key]) = self.match_symbol(*meta_id, model::ORDER_HINT_KEY)? else {
643                    continue;
644                };
645
646                let table::Term::Literal(model::Literal::Nat(key)) = self.get_term(key)? else {
647                    continue;
648                };
649
650                // NOTE: The lookups here are expected to succeed since we only
651                // process the order metadata after we have imported the nodes.
652                let child_node = self.nodes[child_id];
653                let child_optype = self.hugr.get_optype(child_node);
654
655                // Check that the node has order ports.
656                // NOTE: This assumes that a node has an input order port iff it has an output one.
657                if child_optype.other_output_port().is_none() {
658                    return Err(OrderHintError::NoOrderPort(*child_id).into());
659                }
660
661                if order_keys.insert(*key, child_node).is_some() {
662                    return Err(OrderHintError::DuplicateKey(region_id, *key).into());
663                }
664            }
665        }
666
667        // Collect the order hint keys for the input and output nodes
668        for meta_id in region_data.meta {
669            if let Some([key]) = self.match_symbol(*meta_id, model::ORDER_HINT_INPUT_KEY)? {
670                let table::Term::Literal(model::Literal::Nat(key)) = self.get_term(key)? else {
671                    continue;
672                };
673
674                if order_keys.insert(*key, input).is_some() {
675                    return Err(OrderHintError::DuplicateKey(region_id, *key).into());
676                }
677            }
678
679            if let Some([key]) = self.match_symbol(*meta_id, model::ORDER_HINT_OUTPUT_KEY)? {
680                let table::Term::Literal(model::Literal::Nat(key)) = self.get_term(key)? else {
681                    continue;
682                };
683
684                if order_keys.insert(*key, output).is_some() {
685                    return Err(OrderHintError::DuplicateKey(region_id, *key).into());
686                }
687            }
688        }
689
690        // Insert order edges
691        for meta_id in region_data.meta {
692            let Some([a, b]) = self.match_symbol(*meta_id, model::ORDER_HINT_ORDER)? else {
693                continue;
694            };
695
696            let table::Term::Literal(model::Literal::Nat(a)) = self.get_term(a)? else {
697                continue;
698            };
699
700            let table::Term::Literal(model::Literal::Nat(b)) = self.get_term(b)? else {
701                continue;
702            };
703
704            let a = order_keys.get(a).ok_or(OrderHintError::UnknownKey(*a))?;
705            let b = order_keys.get(b).ok_or(OrderHintError::UnknownKey(*b))?;
706
707            // NOTE: The unwrap here must succeed:
708            // - For all ordinary nodes we checked that they have an order port.
709            // - Input and output nodes always have an order port.
710            let a_port = self.hugr.get_optype(*a).other_output_port().unwrap();
711            let b_port = self.hugr.get_optype(*b).other_input_port().unwrap();
712
713            self.hugr.connect(*a, a_port, *b, b_port);
714        }
715
716        Ok(())
717    }
718
719    fn import_adt_and_rest(
720        &mut self,
721        list: table::TermId,
722    ) -> Result<(Vec<TypeRow>, TypeRow), ImportError> {
723        let items = self.import_closed_list(list)?;
724
725        let Some((first, rest)) = items.split_first() else {
726            return Err(error_invalid!("expected list to have at least one element"));
727        };
728
729        let sum_rows: Vec<_> = {
730            let [variants] = self.expect_symbol(*first, model::CORE_ADT)?;
731            self.import_type_rows(variants)?
732        };
733
734        let rest = rest
735            .iter()
736            .map(|term| self.import_type(*term))
737            .collect::<Result<Vec<_>, _>>()?
738            .into();
739
740        Ok((sum_rows, rest))
741    }
742
743    fn import_tail_loop(
744        &mut self,
745        node_id: table::NodeId,
746        parent: Node,
747    ) -> Result<Node, ImportError> {
748        let node_data = self.get_node(node_id)?;
749        debug_assert_eq!(node_data.operation, table::Operation::TailLoop);
750
751        let [region] = node_data.regions else {
752            return Err(error_invalid!(
753                "loop node {} expects a single region",
754                node_id
755            ));
756        };
757
758        let region_data = self.get_region(*region)?;
759
760        let (just_inputs, just_outputs, rest) = (|| {
761            let [_, region_outputs] = self.get_func_type(
762                region_data
763                    .signature
764                    .ok_or_else(|| error_uninferred!("region signature"))?,
765            )?;
766            let (sum_rows, rest) = self.import_adt_and_rest(region_outputs)?;
767
768            if sum_rows.len() != 2 {
769                return Err(error_invalid!(
770                    "loop nodes expect their first target to be an ADT with two variants"
771                ));
772            }
773
774            let mut sum_rows = sum_rows.into_iter();
775            let just_inputs = sum_rows.next().unwrap();
776            let just_outputs = sum_rows.next().unwrap();
777
778            Ok((just_inputs, just_outputs, rest))
779        })()
780        .map_err(|err| error_context!(err, "region signature"))?;
781
782        let optype = OpType::TailLoop(TailLoop {
783            just_inputs,
784            just_outputs,
785            rest,
786        });
787
788        let node = self.make_node(node_id, optype, parent)?;
789
790        self.import_dfg_region(*region, node)?;
791        Ok(node)
792    }
793
794    fn import_conditional(
795        &mut self,
796        node_id: table::NodeId,
797        parent: Node,
798    ) -> Result<Node, ImportError> {
799        let node_data = self.get_node(node_id)?;
800        debug_assert_eq!(node_data.operation, table::Operation::Conditional);
801
802        let (sum_rows, other_inputs, outputs) = (|| {
803            let [inputs, outputs] = self.get_func_type(
804                node_data
805                    .signature
806                    .ok_or_else(|| error_uninferred!("node signature"))?,
807            )?;
808            let (sum_rows, other_inputs) = self.import_adt_and_rest(inputs)?;
809            let outputs = self.import_type_row(outputs)?;
810
811            Ok((sum_rows, other_inputs, outputs))
812        })()
813        .map_err(|err| error_context!(err, "node signature"))?;
814
815        let optype = OpType::Conditional(Conditional {
816            sum_rows,
817            other_inputs,
818            outputs,
819        });
820
821        let node = self.make_node(node_id, optype, parent)?;
822
823        for region in node_data.regions {
824            let region_data = self.get_region(*region)?;
825            let signature = self.import_func_type(
826                region_data
827                    .signature
828                    .ok_or_else(|| error_uninferred!("region signature"))?,
829            )?;
830
831            let case_node = self
832                .hugr
833                .add_node_with_parent(node, OpType::Case(Case { signature }));
834
835            self.import_dfg_region(*region, case_node)?;
836        }
837
838        Ok(node)
839    }
840
841    fn import_cfg_region(
842        &mut self,
843        region: table::RegionId,
844        node: Node,
845    ) -> Result<(), ImportError> {
846        let region_data = self.get_region(region)?;
847
848        if region_data.kind != model::RegionKind::ControlFlow {
849            return Err(error_invalid!("expected cfg region"));
850        }
851
852        let prev_region = self.region_scope;
853        if region_data.scope.is_some() {
854            self.region_scope = region;
855        }
856
857        let region_target_types = (|| {
858            let [_, region_targets] = self.get_ctrl_type(
859                region_data
860                    .signature
861                    .ok_or_else(|| error_uninferred!("region signature"))?,
862            )?;
863
864            self.import_closed_list(region_targets)
865        })()
866        .map_err(|err| error_context!(err, "signature of cfg region with id {}", region))?;
867
868        // Identify the entry node of the control flow region by looking for
869        // a block whose input is linked to the sole source port of the CFG region.
870        let entry_node = 'find_entry: {
871            let [entry_link] = region_data.sources else {
872                return Err(error_invalid!("cfg region expects a single source"));
873            };
874
875            for child in region_data.children {
876                let child_data = self.get_node(*child)?;
877                let is_entry = child_data.inputs.iter().any(|link| link == entry_link);
878
879                if is_entry {
880                    break 'find_entry *child;
881                }
882            }
883
884            // TODO: We should allow for the case in which control flows
885            // directly from the source to the target of the region. This is
886            // currently not allowed in hugr core directly, but may be simulated
887            // by constructing an empty entry block.
888            return Err(error_invalid!("cfg region without entry node"));
889        };
890
891        // The entry node in core control flow regions is identified by being
892        // the first child node of the CFG node. We therefore import the entry node first.
893        self.import_node(entry_node, node)?;
894
895        // Create the exit node for the control flow region. This always needs
896        // to be second in the node list.
897        {
898            let cfg_outputs = {
899                let [target_types] = region_target_types.as_slice() else {
900                    return Err(error_invalid!("cfg region expects a single target"));
901                };
902
903                self.import_type_row(*target_types)?
904            };
905
906            let exit = self
907                .hugr
908                .add_node_with_parent(node, OpType::ExitBlock(ExitBlock { cfg_outputs }));
909            self.record_links(exit, Direction::Incoming, region_data.targets);
910        }
911
912        // Finally we import all other nodes.
913        for child in region_data.children {
914            if *child != entry_node {
915                self.import_node(*child, node)?;
916            }
917        }
918
919        for meta_item in region_data.meta {
920            self.import_node_metadata(node, *meta_item)
921                .map_err(|err| error_context!(err, "node metadata"))?;
922        }
923
924        self.region_scope = prev_region;
925
926        Ok(())
927    }
928
929    fn import_node_block(
930        &mut self,
931        node_id: table::NodeId,
932        parent: Node,
933    ) -> Result<Node, ImportError> {
934        let node_data = self.get_node(node_id)?;
935        debug_assert_eq!(node_data.operation, table::Operation::Block);
936
937        let [region] = node_data.regions else {
938            return Err(error_invalid!("basic block expects a single region"));
939        };
940        let region_data = self.get_region(*region)?;
941        let [inputs, outputs] = self.get_func_type(
942            region_data
943                .signature
944                .ok_or_else(|| error_uninferred!("region signature"))?,
945        )?;
946        let inputs = self.import_type_row(inputs)?;
947        let (sum_rows, other_outputs) = self.import_adt_and_rest(outputs)?;
948
949        let optype = OpType::DataflowBlock(DataflowBlock {
950            inputs,
951            other_outputs,
952            sum_rows,
953        });
954        let node = self.make_node(node_id, optype, parent)?;
955
956        self.import_dfg_region(*region, node).map_err(|err| {
957            error_context!(err, "block body defined by region with id {}", *region)
958        })?;
959        Ok(node)
960    }
961
962    fn import_node_define_func(
963        &mut self,
964        node_id: table::NodeId,
965        symbol: &'a table::Symbol<'a>,
966        node_data: &'a table::Node<'a>,
967        parent: Node,
968    ) -> Result<Node, ImportError> {
969        let visibility = symbol.visibility.clone().ok_or(ImportErrorInner::Invalid(
970            "No visibility for FuncDefn".to_string(),
971        ))?;
972        self.import_poly_func_type(node_id, *symbol, |ctx, signature| {
973            let optype =
974                OpType::FuncDefn(FuncDefn::new_vis(symbol.name, signature, visibility.into()));
975
976            let node = ctx.make_node(node_id, optype, parent)?;
977
978            let [region] = node_data.regions else {
979                return Err(error_invalid!(
980                    "function definition nodes expect a single region"
981                ));
982            };
983
984            ctx.import_dfg_region(*region, node).map_err(|err| {
985                error_context!(err, "function body defined by region with id {}", *region)
986            })?;
987
988            Ok(node)
989        })
990    }
991
992    fn import_node_declare_func(
993        &mut self,
994        node_id: table::NodeId,
995        symbol: &'a table::Symbol<'a>,
996        parent: Node,
997    ) -> Result<Node, ImportError> {
998        let visibility = symbol.visibility.clone().ok_or(ImportErrorInner::Invalid(
999            "No visibility for FuncDecl".to_string(),
1000        ))?;
1001        self.import_poly_func_type(node_id, *symbol, |ctx, signature| {
1002            let optype =
1003                OpType::FuncDecl(FuncDecl::new_vis(symbol.name, signature, visibility.into()));
1004            let node = ctx.make_node(node_id, optype, parent)?;
1005            Ok(node)
1006        })
1007    }
1008
1009    fn import_node_custom(
1010        &mut self,
1011        node_id: table::NodeId,
1012        operation: table::TermId,
1013        node_data: &'a table::Node<'a>,
1014        parent: Node,
1015    ) -> Result<Node, ImportError> {
1016        if let Some([inputs, outputs]) = self.match_symbol(operation, model::CORE_CALL_INDIRECT)? {
1017            let inputs = self.import_type_row(inputs)?;
1018            let outputs = self.import_type_row(outputs)?;
1019            let signature = Signature::new(inputs, outputs);
1020            let optype = OpType::CallIndirect(CallIndirect { signature });
1021            let node = self.make_node(node_id, optype, parent)?;
1022            return Ok(node);
1023        }
1024
1025        if let Some([_, _, func]) = self.match_symbol(operation, model::CORE_CALL)? {
1026            let table::Term::Apply(symbol, args) = self.get_term(func)? else {
1027                return Err(error_invalid!(
1028                    "expected a symbol application to be passed to `{}`",
1029                    model::CORE_CALL
1030                ));
1031            };
1032
1033            let func_sig = self.get_func_signature(*symbol)?;
1034
1035            let type_args = args
1036                .iter()
1037                .map(|term| self.import_term(*term))
1038                .collect::<Result<Vec<TypeArg>, _>>()?;
1039
1040            self.static_edges.push((*symbol, node_id));
1041            let optype = OpType::Call(
1042                Call::try_new(func_sig, type_args).map_err(ImportErrorInner::Signature)?,
1043            );
1044
1045            let node = self.make_node(node_id, optype, parent)?;
1046            return Ok(node);
1047        }
1048
1049        if let Some([_, value]) = self.match_symbol(operation, model::CORE_LOAD_CONST)? {
1050            // If the constant refers directly to a function, import this as the `LoadFunc` operation.
1051            if let table::Term::Apply(symbol, args) = self.get_term(value)? {
1052                let func_node_data = self.get_node(*symbol)?;
1053
1054                if let table::Operation::DefineFunc(_) | table::Operation::DeclareFunc(_) =
1055                    func_node_data.operation
1056                {
1057                    let func_sig = self.get_func_signature(*symbol)?;
1058                    let type_args = args
1059                        .iter()
1060                        .map(|term| self.import_term(*term))
1061                        .collect::<Result<Vec<TypeArg>, _>>()?;
1062
1063                    self.static_edges.push((*symbol, node_id));
1064
1065                    let optype = OpType::LoadFunction(
1066                        LoadFunction::try_new(func_sig, type_args)
1067                            .map_err(ImportErrorInner::Signature)?,
1068                    );
1069
1070                    let node = self.make_node(node_id, optype, parent)?;
1071                    return Ok(node);
1072                }
1073            }
1074
1075            // Otherwise use const nodes
1076            let signature = node_data
1077                .signature
1078                .ok_or_else(|| error_uninferred!("node signature"))?;
1079            let [_, outputs] = self.get_func_type(signature)?;
1080            let outputs = self.import_closed_list(outputs)?;
1081            let output = outputs.first().ok_or_else(|| {
1082                error_invalid!("`{}` expects a single output", model::CORE_LOAD_CONST)
1083            })?;
1084            let datatype = self.import_type(*output)?;
1085
1086            let imported_value = self.import_value(value, *output)?;
1087
1088            let load_const_node = self.make_node(
1089                node_id,
1090                OpType::LoadConstant(LoadConstant {
1091                    datatype: datatype.clone(),
1092                }),
1093                parent,
1094            )?;
1095
1096            let const_node = self
1097                .hugr
1098                .add_node_with_parent(parent, OpType::Const(Const::new(imported_value)));
1099
1100            self.hugr.connect(const_node, 0, load_const_node, 0);
1101
1102            return Ok(load_const_node);
1103        }
1104
1105        if let Some([_, _, tag]) = self.match_symbol(operation, model::CORE_MAKE_ADT)? {
1106            let table::Term::Literal(model::Literal::Nat(tag)) = self.get_term(tag)? else {
1107                return Err(error_invalid!(
1108                    "`{}` expects a nat literal tag",
1109                    model::CORE_MAKE_ADT
1110                ));
1111            };
1112
1113            let signature = node_data
1114                .signature
1115                .ok_or_else(|| error_uninferred!("node signature"))?;
1116            let [_, outputs] = self.get_func_type(signature)?;
1117            let (variants, _) = self.import_adt_and_rest(outputs)?;
1118            let node = self.make_node(
1119                node_id,
1120                OpType::Tag(Tag {
1121                    variants,
1122                    tag: *tag as usize,
1123                }),
1124                parent,
1125            )?;
1126            return Ok(node);
1127        }
1128
1129        let table::Term::Apply(node, params) = self.get_term(operation)? else {
1130            return Err(error_invalid!(
1131                "custom operations expect a symbol application referencing an operation"
1132            ));
1133        };
1134        let name = self.get_symbol_name(*node)?;
1135        let args = params
1136            .iter()
1137            .map(|param| self.import_term(*param))
1138            .collect::<Result<Vec<_>, _>>()?;
1139        let (extension, name) = self.import_custom_name(name)?;
1140        let signature = self.get_node_signature(node_id)?;
1141
1142        // TODO: Currently we do not have the description or any other metadata for
1143        // the custom op. This will improve with declarative extensions being able
1144        // to declare operations as a node, in which case the description will be attached
1145        // to that node as metadata.
1146
1147        let optype = OpType::OpaqueOp(OpaqueOp::new(extension, name, args, signature));
1148        self.make_node(node_id, optype, parent)
1149    }
1150
1151    fn import_node_define_alias(
1152        &mut self,
1153        node_id: table::NodeId,
1154        symbol: &'a table::Symbol<'a>,
1155        value: table::TermId,
1156        parent: Node,
1157    ) -> Result<Node, ImportError> {
1158        if !symbol.params.is_empty() {
1159            return Err(error_unsupported!(
1160                "parameters or constraints in alias definition"
1161            ));
1162        }
1163
1164        let optype = OpType::AliasDefn(AliasDefn {
1165            name: symbol.name.to_smolstr(),
1166            definition: self.import_type(value)?,
1167        });
1168
1169        let node = self.make_node(node_id, optype, parent)?;
1170        Ok(node)
1171    }
1172
1173    fn import_node_declare_alias(
1174        &mut self,
1175        node_id: table::NodeId,
1176        symbol: &'a table::Symbol<'a>,
1177        parent: Node,
1178    ) -> Result<Node, ImportError> {
1179        if !symbol.params.is_empty() {
1180            return Err(error_unsupported!(
1181                "parameters or constraints in alias declaration"
1182            ));
1183        }
1184
1185        let optype = OpType::AliasDecl(AliasDecl {
1186            name: symbol.name.to_smolstr(),
1187            bound: TypeBound::Copyable,
1188        });
1189
1190        let node = self.make_node(node_id, optype, parent)?;
1191        Ok(node)
1192    }
1193
1194    fn import_poly_func_type<RV: MaybeRV, T>(
1195        &mut self,
1196        node: table::NodeId,
1197        symbol: table::Symbol<'a>,
1198        in_scope: impl FnOnce(&mut Self, PolyFuncTypeBase<RV>) -> Result<T, ImportError>,
1199    ) -> Result<T, ImportError> {
1200        (|| {
1201            let mut imported_params = Vec::with_capacity(symbol.params.len());
1202
1203            for (index, param) in symbol.params.iter().enumerate() {
1204                self.local_vars
1205                    .insert(table::VarId(node, index as _), LocalVar::new(param.r#type));
1206            }
1207
1208            for constraint in symbol.constraints {
1209                if let Some([term]) = self.match_symbol(*constraint, model::CORE_NON_LINEAR)? {
1210                    let table::Term::Var(var) = self.get_term(term)? else {
1211                        return Err(error_unsupported!(
1212                            "constraint on term that is not a variable"
1213                        ));
1214                    };
1215
1216                    self.local_vars
1217                        .get_mut(var)
1218                        .ok_or_else(|| error_invalid!("unknown variable {}", var))?
1219                        .bound = TypeBound::Copyable;
1220                } else {
1221                    return Err(error_unsupported!("constraint other than copy or discard"));
1222                }
1223            }
1224
1225            for (index, param) in symbol.params.iter().enumerate() {
1226                // NOTE: `PolyFuncType` only has explicit type parameters at present.
1227                let bound = self.local_vars[&table::VarId(node, index as _)].bound;
1228                imported_params.push(
1229                    self.import_term_with_bound(param.r#type, bound)
1230                        .map_err(|err| error_context!(err, "type of parameter `{}`", param.name))?,
1231                );
1232            }
1233
1234            let body = self.import_func_type::<RV>(symbol.signature)?;
1235            in_scope(self, PolyFuncTypeBase::new(imported_params, body))
1236        })()
1237        .map_err(|err| error_context!(err, "symbol `{}` defined by node {}", symbol.name, node))
1238    }
1239
1240    /// Import a [`Term`] from a term that represents a static type or value.
1241    fn import_term(&mut self, term_id: table::TermId) -> Result<Term, ImportError> {
1242        self.import_term_with_bound(term_id, TypeBound::Linear)
1243    }
1244
1245    fn import_term_with_bound(
1246        &mut self,
1247        term_id: table::TermId,
1248        bound: TypeBound,
1249    ) -> Result<Term, ImportError> {
1250        (|| {
1251            if let Some([]) = self.match_symbol(term_id, model::CORE_STR_TYPE)? {
1252                return Ok(Term::StringType);
1253            }
1254
1255            if let Some([]) = self.match_symbol(term_id, model::CORE_NAT_TYPE)? {
1256                return Ok(Term::max_nat_type());
1257            }
1258
1259            if let Some([]) = self.match_symbol(term_id, model::CORE_BYTES_TYPE)? {
1260                return Ok(Term::BytesType);
1261            }
1262
1263            if let Some([]) = self.match_symbol(term_id, model::CORE_FLOAT_TYPE)? {
1264                return Ok(Term::FloatType);
1265            }
1266
1267            if let Some([]) = self.match_symbol(term_id, model::CORE_TYPE)? {
1268                return Ok(TypeParam::RuntimeType(bound));
1269            }
1270
1271            if let Some([]) = self.match_symbol(term_id, model::CORE_CONSTRAINT)? {
1272                return Err(error_unsupported!("`{}`", model::CORE_CONSTRAINT));
1273            }
1274
1275            if let Some([]) = self.match_symbol(term_id, model::CORE_STATIC)? {
1276                return Ok(Term::StaticType);
1277            }
1278
1279            if let Some([]) = self.match_symbol(term_id, model::CORE_CONST)? {
1280                return Err(error_unsupported!("`{}`", model::CORE_CONST));
1281            }
1282
1283            if let Some([item_type]) = self.match_symbol(term_id, model::CORE_LIST_TYPE)? {
1284                // At present `hugr-model` has no way to express that the item
1285                // type of a list must be copyable. Therefore we import it as `Any`.
1286                let item_type = self
1287                    .import_term(item_type)
1288                    .map_err(|err| error_context!(err, "item type of list type"))?;
1289                return Ok(TypeParam::new_list_type(item_type));
1290            }
1291
1292            if let Some([item_types]) = self.match_symbol(term_id, model::CORE_TUPLE_TYPE)? {
1293                // At present `hugr-model` has no way to express that the item
1294                // types of a tuple must be copyable. Therefore we import it as `Any`.
1295                let item_types = self
1296                    .import_term(item_types)
1297                    .map_err(|err| error_context!(err, "item types of tuple type"))?;
1298                return Ok(TypeParam::new_tuple_type(item_types));
1299            }
1300
1301            match self.get_term(term_id)? {
1302                table::Term::Wildcard => Err(error_uninferred!("wildcard")),
1303
1304                table::Term::Var(var) => {
1305                    let var_info = self
1306                        .local_vars
1307                        .get(var)
1308                        .ok_or_else(|| error_invalid!("unknown variable {}", var))?;
1309                    let decl = self.import_term_with_bound(var_info.r#type, var_info.bound)?;
1310                    Ok(Term::new_var_use(var.1 as _, decl))
1311                }
1312
1313                table::Term::List(parts) => {
1314                    // PERFORMANCE: Can we do this without the additional allocation?
1315                    let parts: Vec<_> = parts
1316                        .iter()
1317                        .map(|part| self.import_seq_part(part))
1318                        .collect::<Result<_, _>>()
1319                        .map_err(|err| error_context!(err, "list parts"))?;
1320                    Ok(TypeArg::new_list_from_parts(parts))
1321                }
1322
1323                table::Term::Tuple(parts) => {
1324                    // PERFORMANCE: Can we do this without the additional allocation?
1325                    let parts: Vec<_> = parts
1326                        .iter()
1327                        .map(|part| self.import_seq_part(part))
1328                        .try_collect()
1329                        .map_err(|err| error_context!(err, "tuple parts"))?;
1330                    Ok(TypeArg::new_tuple_from_parts(parts))
1331                }
1332
1333                table::Term::Literal(model::Literal::Str(value)) => {
1334                    Ok(Term::String(value.to_string()))
1335                }
1336
1337                table::Term::Literal(model::Literal::Nat(value)) => Ok(Term::BoundedNat(*value)),
1338
1339                table::Term::Literal(model::Literal::Bytes(value)) => {
1340                    Ok(Term::Bytes(value.clone()))
1341                }
1342                table::Term::Literal(model::Literal::Float(value)) => Ok(Term::Float(*value)),
1343                table::Term::Func { .. } => Err(error_unsupported!("function constant")),
1344
1345                table::Term::Apply { .. } => {
1346                    let ty: Type = self.import_type(term_id)?;
1347                    Ok(ty.into())
1348                }
1349            }
1350        })()
1351        .map_err(|err| error_context!(err, "term {}", term_id))
1352    }
1353
1354    fn import_seq_part(
1355        &mut self,
1356        seq_part: &'a table::SeqPart,
1357    ) -> Result<SeqPart<TypeArg>, ImportError> {
1358        Ok(match seq_part {
1359            table::SeqPart::Item(term_id) => SeqPart::Item(self.import_term(*term_id)?),
1360            table::SeqPart::Splice(term_id) => SeqPart::Splice(self.import_term(*term_id)?),
1361        })
1362    }
1363
1364    /// Import a `Type` from a term that represents a runtime type.
1365    fn import_type<RV: MaybeRV>(
1366        &mut self,
1367        term_id: table::TermId,
1368    ) -> Result<TypeBase<RV>, ImportError> {
1369        (|| {
1370            if let Some([_, _]) = self.match_symbol(term_id, model::CORE_FN)? {
1371                let func_type = self.import_func_type::<RowVariable>(term_id)?;
1372                return Ok(TypeBase::new_function(func_type));
1373            }
1374
1375            if let Some([variants]) = self.match_symbol(term_id, model::CORE_ADT)? {
1376                let variants = (|| {
1377                    self.import_closed_list(variants)?
1378                        .iter()
1379                        .map(|variant| self.import_type_row::<RowVariable>(*variant))
1380                        .collect::<Result<Vec<_>, _>>()
1381                })()
1382                .map_err(|err| error_context!(err, "adt variants"))?;
1383
1384                return Ok(TypeBase::new_sum(variants));
1385            }
1386
1387            match self.get_term(term_id)? {
1388                table::Term::Wildcard => Err(error_uninferred!("wildcard")),
1389
1390                table::Term::Apply(symbol, args) => {
1391                    let name = self.get_symbol_name(*symbol)?;
1392
1393                    let args = args
1394                        .iter()
1395                        .map(|arg| self.import_term(*arg))
1396                        .collect::<Result<Vec<_>, _>>()
1397                        .map_err(|err| {
1398                            error_context!(err, "type argument of custom type `{}`", name)
1399                        })?;
1400
1401                    let (extension, id) = self.import_custom_name(name)?;
1402
1403                    let extension_ref =
1404                        self.extensions
1405                            .get(&extension)
1406                            .ok_or_else(|| ExtensionError::Missing {
1407                                missing_ext: extension.clone(),
1408                                available: self.extensions.ids().cloned().collect(),
1409                            })?;
1410
1411                    let ext_type =
1412                        extension_ref
1413                            .get_type(&id)
1414                            .ok_or_else(|| ExtensionError::MissingType {
1415                                ext: extension.clone(),
1416                                name: id.clone(),
1417                            })?;
1418
1419                    let bound = ext_type.bound(&args);
1420
1421                    Ok(TypeBase::new_extension(CustomType::new(
1422                        id,
1423                        args,
1424                        extension,
1425                        bound,
1426                        &Arc::downgrade(extension_ref),
1427                    )))
1428                }
1429
1430                table::Term::Var(var @ table::VarId(_, index)) => {
1431                    let local_var = self
1432                        .local_vars
1433                        .get(var)
1434                        .ok_or(error_invalid!("unknown var {}", var))?;
1435                    Ok(TypeBase::new_var_use(*index as _, local_var.bound))
1436                }
1437
1438                // The following terms are not runtime types, but the core `Type` only contains runtime types.
1439                // We therefore report a type error here.
1440                table::Term::List { .. }
1441                | table::Term::Tuple { .. }
1442                | table::Term::Literal(_)
1443                | table::Term::Func { .. } => Err(error_invalid!("expected a runtime type")),
1444            }
1445        })()
1446        .map_err(|err| error_context!(err, "term {} as `Type`", term_id))
1447    }
1448
1449    fn get_func_type(&mut self, term_id: table::TermId) -> Result<[table::TermId; 2], ImportError> {
1450        self.match_symbol(term_id, model::CORE_FN)?
1451            .ok_or(error_invalid!("expected a function type"))
1452    }
1453
1454    fn get_ctrl_type(&mut self, term_id: table::TermId) -> Result<[table::TermId; 2], ImportError> {
1455        self.match_symbol(term_id, model::CORE_CTRL)?
1456            .ok_or(error_invalid!("expected a control type"))
1457    }
1458
1459    fn import_func_type<RV: MaybeRV>(
1460        &mut self,
1461        term_id: table::TermId,
1462    ) -> Result<FuncTypeBase<RV>, ImportError> {
1463        (|| {
1464            let [inputs, outputs] = self.get_func_type(term_id)?;
1465            let inputs = self
1466                .import_type_row(inputs)
1467                .map_err(|err| error_context!(err, "function inputs"))?;
1468            let outputs = self
1469                .import_type_row(outputs)
1470                .map_err(|err| error_context!(err, "function outputs"))?;
1471            Ok(FuncTypeBase::new(inputs, outputs))
1472        })()
1473        .map_err(|err| error_context!(err, "function type"))
1474    }
1475
1476    fn import_closed_list(
1477        &mut self,
1478        term_id: table::TermId,
1479    ) -> Result<Vec<table::TermId>, ImportError> {
1480        fn import_into(
1481            ctx: &mut Context,
1482            term_id: table::TermId,
1483            types: &mut Vec<table::TermId>,
1484        ) -> Result<(), ImportError> {
1485            match ctx.get_term(term_id)? {
1486                table::Term::List(parts) => {
1487                    types.reserve(parts.len());
1488
1489                    for part in *parts {
1490                        match part {
1491                            table::SeqPart::Item(term_id) => {
1492                                types.push(*term_id);
1493                            }
1494                            table::SeqPart::Splice(term_id) => {
1495                                import_into(ctx, *term_id, types)?;
1496                            }
1497                        }
1498                    }
1499                }
1500                _ => return Err(error_invalid!("expected a closed list")),
1501            }
1502
1503            Ok(())
1504        }
1505
1506        let mut types = Vec::new();
1507        import_into(self, term_id, &mut types)?;
1508        Ok(types)
1509    }
1510
1511    fn import_closed_tuple(
1512        &mut self,
1513        term_id: table::TermId,
1514    ) -> Result<Vec<table::TermId>, ImportError> {
1515        fn import_into(
1516            ctx: &mut Context,
1517            term_id: table::TermId,
1518            types: &mut Vec<table::TermId>,
1519        ) -> Result<(), ImportError> {
1520            match ctx.get_term(term_id)? {
1521                table::Term::Tuple(parts) => {
1522                    types.reserve(parts.len());
1523
1524                    for part in *parts {
1525                        match part {
1526                            table::SeqPart::Item(term_id) => {
1527                                types.push(*term_id);
1528                            }
1529                            table::SeqPart::Splice(term_id) => {
1530                                import_into(ctx, *term_id, types)?;
1531                            }
1532                        }
1533                    }
1534                }
1535                _ => return Err(error_invalid!("expected a closed tuple")),
1536            }
1537
1538            Ok(())
1539        }
1540
1541        let mut types = Vec::new();
1542        import_into(self, term_id, &mut types)?;
1543        Ok(types)
1544    }
1545
1546    fn import_type_rows<RV: MaybeRV>(
1547        &mut self,
1548        term_id: table::TermId,
1549    ) -> Result<Vec<TypeRowBase<RV>>, ImportError> {
1550        self.import_closed_list(term_id)?
1551            .into_iter()
1552            .map(|term_id| self.import_type_row::<RV>(term_id))
1553            .collect()
1554    }
1555
1556    fn import_type_row<RV: MaybeRV>(
1557        &mut self,
1558        term_id: table::TermId,
1559    ) -> Result<TypeRowBase<RV>, ImportError> {
1560        fn import_into<RV: MaybeRV>(
1561            ctx: &mut Context,
1562            term_id: table::TermId,
1563            types: &mut Vec<TypeBase<RV>>,
1564        ) -> Result<(), ImportError> {
1565            match ctx.get_term(term_id)? {
1566                table::Term::List(parts) => {
1567                    types.reserve(parts.len());
1568
1569                    for item in *parts {
1570                        match item {
1571                            table::SeqPart::Item(term_id) => {
1572                                types.push(ctx.import_type::<RV>(*term_id)?);
1573                            }
1574                            table::SeqPart::Splice(term_id) => {
1575                                import_into(ctx, *term_id, types)?;
1576                            }
1577                        }
1578                    }
1579                }
1580                table::Term::Var(table::VarId(_, index)) => {
1581                    let var = RV::try_from_rv(RowVariable(*index as _, TypeBound::Linear))
1582                        .map_err(|_| error_invalid!("expected a closed list"))?;
1583                    types.push(TypeBase::new(TypeEnum::RowVar(var)));
1584                }
1585                _ => return Err(error_invalid!("expected a list")),
1586            }
1587
1588            Ok(())
1589        }
1590
1591        let mut types = Vec::new();
1592        import_into(self, term_id, &mut types)?;
1593        Ok(types.into())
1594    }
1595
1596    fn import_custom_name(
1597        &mut self,
1598        symbol: &'a str,
1599    ) -> Result<(ExtensionId, SmolStr), ImportError> {
1600        use std::collections::hash_map::Entry;
1601        match self.custom_name_cache.entry(symbol) {
1602            Entry::Occupied(occupied_entry) => Ok(occupied_entry.get().clone()),
1603            Entry::Vacant(vacant_entry) => {
1604                let qualified_name = ExtensionId::new(symbol)
1605                    .map_err(|_| error_invalid!("`{}` is not a valid symbol name", symbol))?;
1606
1607                let (extension, id) = qualified_name
1608                    .split_last()
1609                    .ok_or_else(|| error_invalid!("`{}` is not a valid symbol name", symbol))?;
1610
1611                vacant_entry.insert((extension.clone(), id.clone()));
1612                Ok((extension, id))
1613            }
1614        }
1615    }
1616
1617    fn import_value(
1618        &mut self,
1619        term_id: table::TermId,
1620        type_id: table::TermId,
1621    ) -> Result<Value, ImportError> {
1622        let term_data = self.get_term(term_id)?;
1623
1624        // NOTE: We have special cased arrays, integers, and floats for now.
1625        // TODO: Allow arbitrary extension values to be imported from terms.
1626
1627        if let Some([runtime_type, json]) = self.match_symbol(term_id, model::COMPAT_CONST_JSON)? {
1628            let table::Term::Literal(model::Literal::Str(json)) = self.get_term(json)? else {
1629                return Err(error_invalid!(
1630                    "`{}` expects a string literal",
1631                    model::COMPAT_CONST_JSON
1632                ));
1633            };
1634
1635            // We attempt to deserialize as the custom const directly.
1636            // This might fail due to the custom const struct not being included when
1637            // this code was compiled; in that case, we fall back to the serialized form.
1638            let value: Option<Box<dyn CustomConst>> = serde_json::from_str(json).ok();
1639
1640            if let Some(value) = value {
1641                let opaque_value = OpaqueValue::from(value);
1642                return Ok(Value::Extension { e: opaque_value });
1643            } else {
1644                let runtime_type = self.import_type(runtime_type)?;
1645                let value: serde_json::Value = serde_json::from_str(json).map_err(|_| {
1646                    error_invalid!(
1647                        "unable to parse JSON string for `{}`",
1648                        model::COMPAT_CONST_JSON
1649                    )
1650                })?;
1651                let custom_const = CustomSerialized::new(runtime_type, value);
1652                let opaque_value = OpaqueValue::new(custom_const);
1653                return Ok(Value::Extension { e: opaque_value });
1654            }
1655        }
1656
1657        if let Some([_, element_type_term, contents]) =
1658            self.match_symbol(term_id, ArrayValue::CTR_NAME)?
1659        {
1660            let element_type = self.import_type(element_type_term)?;
1661            let contents = self.import_closed_list(contents)?;
1662            let contents = contents
1663                .iter()
1664                .map(|item| self.import_value(*item, element_type_term))
1665                .collect::<Result<Vec<_>, _>>()?;
1666            return Ok(ArrayValue::new(element_type, contents).into());
1667        }
1668
1669        if let Some([bitwidth, value]) = self.match_symbol(term_id, ConstInt::CTR_NAME)? {
1670            let bitwidth = {
1671                let table::Term::Literal(model::Literal::Nat(bitwidth)) =
1672                    self.get_term(bitwidth)?
1673                else {
1674                    return Err(error_invalid!(
1675                        "`{}` expects a nat literal in its `bitwidth` argument",
1676                        ConstInt::CTR_NAME
1677                    ));
1678                };
1679                if *bitwidth > 6 {
1680                    return Err(error_invalid!(
1681                        "`{}` expects the bitwidth to be at most 6, got {}",
1682                        ConstInt::CTR_NAME,
1683                        bitwidth
1684                    ));
1685                }
1686                *bitwidth as u8
1687            };
1688
1689            let value = {
1690                let table::Term::Literal(model::Literal::Nat(value)) = self.get_term(value)? else {
1691                    return Err(error_invalid!(
1692                        "`{}` expects a nat literal value",
1693                        ConstInt::CTR_NAME
1694                    ));
1695                };
1696                *value
1697            };
1698
1699            return Ok(ConstInt::new_u(bitwidth, value)
1700                .map_err(|_| error_invalid!("failed to create int constant"))?
1701                .into());
1702        }
1703
1704        if let Some([value]) = self.match_symbol(term_id, ConstF64::CTR_NAME)? {
1705            let table::Term::Literal(model::Literal::Float(value)) = self.get_term(value)? else {
1706                return Err(error_invalid!(
1707                    "`{}` expects a float literal value",
1708                    ConstF64::CTR_NAME
1709                ));
1710            };
1711
1712            return Ok(ConstF64::new(value.into_inner()).into());
1713        }
1714
1715        if let Some([_, _, tag, values]) = self.match_symbol(term_id, model::CORE_CONST_ADT)? {
1716            let [variants] = self.expect_symbol(type_id, model::CORE_ADT)?;
1717            let values = self.import_closed_tuple(values)?;
1718            let variants = self.import_closed_list(variants)?;
1719
1720            let table::Term::Literal(model::Literal::Nat(tag)) = self.get_term(tag)? else {
1721                return Err(error_invalid!(
1722                    "`{}` expects a nat literal tag",
1723                    model::CORE_ADT
1724                ));
1725            };
1726
1727            let variant = variants.get(*tag as usize).ok_or(error_invalid!(
1728                "the tag of a `{}` must be a valid index into the list of variants",
1729                model::CORE_CONST_ADT
1730            ))?;
1731
1732            let variant = self.import_closed_list(*variant)?;
1733
1734            let items = values
1735                .iter()
1736                .zip(variant.iter())
1737                .map(|(value, typ)| self.import_value(*value, *typ))
1738                .collect::<Result<Vec<_>, _>>()?;
1739
1740            let typ = {
1741                // TODO: Import as a `SumType` directly and avoid the copy.
1742                let typ: Type = self.import_type(type_id)?;
1743                match typ.as_type_enum() {
1744                    TypeEnum::Sum(sum) => sum.clone(),
1745                    _ => unreachable!(),
1746                }
1747            };
1748
1749            return Ok(Value::sum(*tag as _, items, typ).unwrap());
1750        }
1751
1752        match term_data {
1753            table::Term::Wildcard => Err(error_uninferred!("wildcard")),
1754            table::Term::Var(_) => Err(error_unsupported!("constant value containing a variable")),
1755
1756            table::Term::Apply(symbol, _) => {
1757                let symbol_name = self.get_symbol_name(*symbol)?;
1758                Err(error_unsupported!(
1759                    "unknown custom constant value `{}`",
1760                    symbol_name
1761                ))
1762                // TODO: This should ultimately include the following cases:
1763                // - function definitions
1764                // - custom constructors for values
1765            }
1766
1767            table::Term::List { .. } | table::Term::Tuple(_) | table::Term::Literal(_) => {
1768                Err(error_invalid!("expected constant"))
1769            }
1770
1771            table::Term::Func { .. } => Err(error_unsupported!("constant function value")),
1772        }
1773    }
1774
1775    fn match_symbol<const N: usize>(
1776        &self,
1777        term_id: table::TermId,
1778        name: &str,
1779    ) -> Result<Option<[table::TermId; N]>, ImportError> {
1780        let term = self.get_term(term_id)?;
1781
1782        // TODO: Follow alias chains?
1783
1784        let table::Term::Apply(symbol, args) = term else {
1785            return Ok(None);
1786        };
1787
1788        if name != self.get_symbol_name(*symbol)? {
1789            return Ok(None);
1790        }
1791
1792        // We allow the match even if the symbol is applied to fewer arguments
1793        // than parameters. In that case, the arguments are padded with wildcards
1794        // at the beginning.
1795        if args.len() > N {
1796            return Ok(None);
1797        }
1798
1799        let result = std::array::from_fn(|i| {
1800            (i + args.len())
1801                .checked_sub(N)
1802                .map(|i| args[i])
1803                .unwrap_or_default()
1804        });
1805
1806        Ok(Some(result))
1807    }
1808
1809    fn expect_symbol<const N: usize>(
1810        &self,
1811        term_id: table::TermId,
1812        name: &str,
1813    ) -> Result<[table::TermId; N], ImportError> {
1814        self.match_symbol(term_id, name)?.ok_or(error_invalid!(
1815            "expected symbol `{}` with arity {}",
1816            name,
1817            N
1818        ))
1819    }
1820}
1821
1822/// Information about a local variable.
1823#[derive(Debug, Clone, Copy)]
1824struct LocalVar {
1825    /// The type of the variable.
1826    r#type: table::TermId,
1827    /// The type bound of the variable.
1828    bound: TypeBound,
1829}
1830
1831impl LocalVar {
1832    pub fn new(r#type: table::TermId) -> Self {
1833        Self {
1834            r#type,
1835            bound: TypeBound::Linear,
1836        }
1837    }
1838}